aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorGravatar Jason Gross <jagro@google.com>2018-07-25 13:25:25 -0400
committerGravatar Jason Gross <jasongross9@gmail.com>2018-07-25 14:52:37 -0400
commitfa9c7dbe23d632d075a17b9f624adff1f1a923fa (patch)
tree58d061a06a6a0c1254fe3d7e752dc731cf2068e9
parentf3fdfecccafccf3828604072f52d541cf9324e67 (diff)
Improve rewriter speed
Andres and I met yesterday, and discovered that there's a source of non-linear complexity in the rewriter which is not type casts. In adding side-conditions to the rewrite rules (which are not discussed in the pattern-matching compilation paper), I represented them by allowing rewrite rules to fail. So, for example, # + x ~~> x (when # == 0) is represented as # + x ~~> if (# =? 0) then Some x else None In the case that a rewrite rule fails, we need to try all other rewrite rules that might still apply. However, doing this in the naive-CPS way leads to non-linear blowup, because wildcard rewrite rules get duplicated in the failure branches. (This is similar to the issue that `match x with "some string" => true | _ => false end%string` will generate a large number of "false" branches, and duplicate "false" across all of them, rather than having a single default case.) For example, if we had the rewrite rules # + # ~~> literal sum x + (-y) ~~> x - y (-x) + y ~~> y - x then the compiled code would look like fun x y => if x is a literal then if y is a literal then literal sum else if y is an opp then x - y else x + y else if y is an opp then x - y else if x is an opp then y - x else x + y where we actually want the code fun x y => if x is a literal then if y is a literal then return (literal sum); if y is an opp then return (x - y); if x is an opp then return (y - x); return (x + y) in the sequence+return monad. i.e., we want to not duplicate the "if y is an opp" code multiple times. I think the solution to this is to have the discrimination tree evaluator return an option, and to have the function that computes the discrimination tree not duplicate rewrite rules among different cases. Note that this leads to slightly inefficient matching sometimes: when two rules with the same structure are separated by a rule with a wildcard instead of structure, we will now try to match on the structure twice. It might be useful to be able to denote that some rewrite rules can be commuted. Also, we remove `Module pident := pattern.ident.` to work around https://github.com/coq/coq/issues/8152, now that we're no longer unfolding all of the `pattern.ident` stuff (because not unfolding it gives slightly nicer output code). After | File Name | Before || Change | % Change --------------------------------------------------------------------------------------------------------------------- 17m41.82s | Total | 29m58.10s || -12m16.27s | -40.94% --------------------------------------------------------------------------------------------------------------------- 1m08.84s | Experiments/NewPipeline/Rewriter | 5m16.96s || -4m08.11s | -78.28% 0m37.50s | Experiments/NewPipeline/ExtractionHaskell/word_by_word_montgomery | 2m03.59s || -1m26.09s | -69.65% 0m17.35s | Experiments/NewPipeline/ExtractionHaskell/saturated_solinas | 1m41.56s || -1m24.21s | -82.91% 0m24.10s | Experiments/NewPipeline/ExtractionHaskell/unsaturated_solinas | 1m43.33s || -1m19.22s | -76.67% 0m38.44s | Experiments/NewPipeline/ExtractionOCaml/word_by_word_montgomery | 1m28.64s || -0m50.20s | -56.63% 0m20.86s | Experiments/NewPipeline/ExtractionOCaml/unsaturated_solinas | 1m06.91s || -0m46.04s | -68.82% 0m11.46s | Experiments/NewPipeline/ExtractionOCaml/saturated_solinas | 0m55.64s || -0m44.17s | -79.40% 0m04.17s | Experiments/NewPipeline/ExtractionOCaml/saturated_solinas.ml | 0m22.22s || -0m18.04s | -81.23% 1m31.73s | Experiments/NewPipeline/Toplevel2 | 1m49.50s || -0m17.76s | -16.22% 0m06.12s | Experiments/NewPipeline/ExtractionOCaml/unsaturated_solinas.ml | 0m23.22s || -0m17.09s | -73.64% 0m08.69s | Experiments/NewPipeline/ExtractionOCaml/word_by_word_montgomery.ml | 0m22.12s || -0m13.43s | -60.71% 0m05.80s | Experiments/NewPipeline/ExtractionHaskell/word_by_word_montgomery.hs | 0m15.32s || -0m09.51s | -62.14% 0m04.10s | Experiments/NewPipeline/ExtractionHaskell/unsaturated_solinas.hs | 0m13.20s || -0m09.09s | -68.93% 0m03.30s | Experiments/NewPipeline/ExtractionHaskell/saturated_solinas.hs | 0m13.02s || -0m09.71s | -74.65% 5m47.58s | Experiments/NewPipeline/SlowPrimeSynthesisExamples | 5m52.86s || -0m05.28s | -1.49% 4m02.15s | Experiments/NewPipeline/Toplevel1 | 4m04.75s || -0m02.59s | -1.06% 0m11.40s | p384_64.c | 0m09.04s || +0m02.36s | +26.10% 0m34.20s | p521_64.c | 0m32.80s || +0m01.40s | +4.26% 0m23.40s | p384_32.c | 0m21.95s || +0m01.44s | +6.60% 0m38.90s | p521_32.c | 0m39.34s || -0m00.44s | -1.11% 0m03.40s | p256_32.c | 0m03.63s || -0m00.23s | -6.33% 0m03.24s | secp256k1_32.c | 0m03.37s || -0m00.12s | -3.85% 0m02.14s | curve25519_32.c | 0m02.00s || +0m00.14s | +7.00% 0m02.13s | p256_64.c | 0m01.72s || +0m00.40s | +23.83% 0m01.82s | p224_32.c | 0m01.82s || +0m00.00s | +0.00% 0m01.44s | p224_64.c | 0m01.48s || -0m00.04s | -2.70% 0m01.40s | secp256k1_64.c | 0m01.97s || -0m00.57s | -28.93% 0m01.35s | curve25519_64.c | 0m01.34s || +0m00.01s | +0.74% 0m01.34s | Experiments/NewPipeline/CLI | 0m01.32s || +0m00.02s | +1.51% 0m01.20s | Experiments/NewPipeline/StandaloneHaskellMain | 0m01.16s || +0m00.04s | +3.44% 0m01.20s | Experiments/NewPipeline/StandaloneOCamlMain | 0m01.25s || -0m00.05s | -4.00% 0m01.08s | Experiments/NewPipeline/CompilersTestCases | 0m01.07s || +0m00.01s | +0.93%
-rw-r--r--src/Experiments/NewPipeline/Rewriter.v481
-rw-r--r--src/Experiments/NewPipeline/StandaloneHaskellMain.v2
-rw-r--r--src/Experiments/NewPipeline/StandaloneOCamlMain.v2
-rw-r--r--src/Experiments/NewPipeline/arith_rewrite_head.out18099
-rw-r--r--src/Experiments/NewPipeline/fancy_rewrite_head.out36419
-rw-r--r--src/Experiments/NewPipeline/nbe_rewrite_head.out6766
6 files changed, 6082 insertions, 55687 deletions
diff --git a/src/Experiments/NewPipeline/Rewriter.v b/src/Experiments/NewPipeline/Rewriter.v
index f9f6e22ac..794dede40 100644
--- a/src/Experiments/NewPipeline/Rewriter.v
+++ b/src/Experiments/NewPipeline/Rewriter.v
@@ -8,6 +8,7 @@ Require Crypto.Util.PrimitiveHList.
Require Import Crypto.Experiments.NewPipeline.Language.
Require Import Crypto.Experiments.NewPipeline.UnderLets.
Require Import Crypto.Experiments.NewPipeline.GENERATEDIdentifiersWithoutTypes.
+Require Import Crypto.Util.LetIn.
Require Import Crypto.Util.Notations.
Import ListNotations. Local Open Scope bool_scope. Local Open Scope Z_scope.
@@ -415,66 +416,79 @@ Module Compilers.
| _, _ => None
end.
- Fixpoint eval_decision_tree {T} (ctx : list rawexpr) (d : decision_tree) (cont : option nat -> list rawexpr -> option (unit -> T) -> T) {struct d} : T
+ Definition option_bind' {A B} := @Option.bind A B. (* for help with rewriting *)
+
+ Fixpoint eval_decision_tree {T} (ctx : list rawexpr) (d : decision_tree) (cont : nat -> list rawexpr -> option T) {struct d} : option T
:= match d with
| TryLeaf k onfailure
- => cont (Some k) ctx
- (Some (fun 'tt => @eval_decision_tree T ctx onfailure cont))
- | Failure => cont None ctx None
+ => let res := cont k ctx in
+ match onfailure with
+ | Failure => res
+ | _ => res ;; (@eval_decision_tree T ctx onfailure cont)
+ end
+ | Failure => None
| Switch icases app_case default_case
=> match ctx with
- | nil => cont None ctx None
+ | nil => None
| ctx0 :: ctx'
- => let default _ := @eval_decision_tree T ctx default_case cont in
- reveal_rawexpr_cps
- ctx0 _
- (fun ctx0'
- => match ctx0' with
- | rIdent t idc t' alt
- => fold_right
- (fun '(pidc, icase) default 'tt
- => match invert_bind_args _ idc pidc with
- | Some args
- => @eval_decision_tree
- T ctx' icase
- (fun k ctx''
- => cont k (rIdent (pident_to_typed pidc args) alt :: ctx''))
- | None => default tt
- end)
- default
- icases
- tt
- | rApp f x t alt
- => match app_case with
- | Some app_case
- => @eval_decision_tree
- T (f :: x :: ctx') app_case
- (fun k ctx''
- => match ctx'' with
- | f' :: x' :: ctx'''
- => cont k (rApp f' x' alt :: ctx''')
- | _ => cont None ctx
- end)
- | None => default tt
- end
- | rExpr t e
- | rValue t e
- => default tt
- end)
+ => let res
+ := reveal_rawexpr_cps
+ ctx0 _
+ (fun ctx0'
+ => match ctx0' with
+ | rIdent t idc t' alt
+ => fold_right
+ (fun '(pidc, icase) rest
+ => let res
+ := option_bind'
+ (invert_bind_args _ idc pidc)
+ (fun args
+ => @eval_decision_tree
+ T ctx' icase
+ (fun k ctx''
+ => cont k (rIdent (pident_to_typed pidc args) alt :: ctx''))) in
+ match rest with
+ | None => Some res
+ | Some rest => Some (res ;; rest)
+ end)
+ None
+ icases;;;
+ None
+ | rApp f x t alt
+ => match app_case with
+ | Some app_case
+ => @eval_decision_tree
+ T (f :: x :: ctx') app_case
+ (fun k ctx''
+ => match ctx'' with
+ | f' :: x' :: ctx'''
+ => cont k (rApp f' x' alt :: ctx''')
+ | _ => None
+ end)
+ | None => None
+ end
+ | rExpr t e
+ | rValue t e
+ => None
+ end) in
+ match default_case with
+ | Failure => res
+ | _ => res ;; (@eval_decision_tree T ctx default_case cont)
+ end
end
| Swap i d'
=> match swap_list 0 i ctx with
- | Some ctx'
- => @eval_decision_tree
- T ctx' d'
- (fun k ctx''
- => match swap_list 0 i ctx'' with
- | Some ctx''' => cont k ctx'''
- | None => cont None ctx
- end)
- | None => cont None ctx None
- end
- end.
+ | Some ctx'
+ => @eval_decision_tree
+ T ctx' d'
+ (fun k ctx''
+ => match swap_list 0 i ctx'' with
+ | Some ctx''' => cont k ctx'''
+ | None => None
+ end)
+ | None => None
+ end
+ end%option.
Local Notation opt_anyexprP ivar
:= (fun should_do_again : bool => UnderLets (@AnyExpr.anyexpr base.type ident (if should_do_again then ivar else var)))
@@ -488,6 +502,8 @@ Module Compilers.
Definition rewrite_rulesT
:= (list rewrite_ruleT).
+ Definition ERROR_BAD_REWRITE_RULE {t} (pat : pattern) (value : expr t) : expr t. exact value. Qed.
+
Definition eval_rewrite_rules
(do_again : forall t : base.type, @expr.expr base.type ident value t -> UnderLets (expr t))
(maybe_do_again
@@ -499,43 +515,33 @@ Module Compilers.
(rew : rewrite_rulesT)
(e : rawexpr)
: UnderLets (expr (type_of_rawexpr e))
- := eval_decision_tree
- (e::nil) d
- (fun k ctx default_on_rewrite_failure
- => match k, ctx return UnderLets (expr (type_of_rawexpr e)) with
- | Some k', e'::nil
- => match nth_error rew k' return UnderLets (expr (type_of_rawexpr e)) with
- | Some (existT p f)
- => bind_data_cps
+ := let defaulte := expr_of_rawexpr e in
+ (eval_decision_tree
+ (e::nil) d
+ (fun k ctx
+ => match ctx return option (UnderLets (expr (type_of_rawexpr e))) with
+ | e'::nil
+ => (pf <- nth_error rew k;
+ let 'existT p f := pf in
+ bind_data_cps
e' p _
(fun v
- => match v with
- | Some v
- => f v _
- (fun fv
- => match fv return UnderLets (expr (type_of_rawexpr e)) with
- | Some (existT should_do_again fv)
- => (fv <-- fv;
- fv <-- maybe_do_again should_do_again _ fv;
- type.try_transport_cps
- base.try_make_transport_cps _ _ _ fv _
- (fun fv'
- => match fv', default_on_rewrite_failure with
- | Some fv'', _ => UnderLets.Base fv''
- | None, Some default => default tt
- | None, None => UnderLets.Base (expr_of_rawexpr e)
- end))%under_lets
- | None => match default_on_rewrite_failure with
- | Some default => default tt
- | None => UnderLets.Base (expr_of_rawexpr e)
- end
- end)
- | None => UnderLets.Base (expr_of_rawexpr e)
- end)
- | None => UnderLets.Base (expr_of_rawexpr e)
- end
- | _, _ => UnderLets.Base (expr_of_rawexpr e)
- end).
+ => v <- v;
+ f v _
+ (fun fv
+ => fv <- fv;
+ let 'existT should_do_again fv := fv in
+ Some
+ (fv <-- fv;
+ fv <-- maybe_do_again should_do_again _ fv;
+ type.try_transport_cps
+ base.try_make_transport_cps _ _ _ fv _
+ (fun fv'
+ => UnderLets.Base
+ (fv' ;;; (ERROR_BAD_REWRITE_RULE p defaulte))))%under_lets)))%option
+ | _ => None
+ end);;;
+ (UnderLets.Base defaulte))%option.
Local Notation enumerate ls
:= (List.combine (List.seq 0 (List.length ls)) ls).
@@ -558,12 +564,20 @@ Module Compilers.
end)
(enumerate p).
+ Definition starts_with_wildcard : nat * list pattern -> bool
+ := fun '(_, p) => match p with
+ | pattern.Wildcard _::_ => true
+ | _ => false
+ end.
+
+ Definition not_starts_with_wildcard : nat * list pattern -> bool
+ := fun p => negb (starts_with_wildcard p).
+
Definition filter_pattern_wildcard (p : list (nat * list pattern)) : list (nat * list pattern)
- := filter (fun '(_, p) => match p with
- | pattern.Wildcard _::_ => true
- | _ => false
- end)
- p.
+ := filter starts_with_wildcard p.
+
+ Definition split_at_first_pattern_wildcard (p : list (nat * list pattern)) : list (nat * list pattern) * list (nat * list pattern)
+ := (take_while not_starts_with_wildcard p, drop_while not_starts_with_wildcard p).
Fixpoint get_unique_pattern_ident' (p : list (nat * list pattern)) (so_far : list pident) : list pident
:= match p with
@@ -592,25 +606,23 @@ Module Compilers.
end)
p.
- Definition refine_pattern_app (p : nat * list pattern) : option (nat * list pattern)
+ Definition filter_pattern_app (p : nat * list pattern) : option (nat * list pattern)
:= match p with
- | (n, pattern.Wildcard d::ps)
- => Some (n, (??{?? -> d} :: ?? :: ps)%list%pattern)
| (n, pattern.App f x :: ps)
=> Some (n, f :: x :: ps)
| (_, pattern.Ident _::_)
+ | (_, pattern.Wildcard _::_)
| (_, nil)
=> None
end.
- Definition refine_pattern_pident (pidc : pident) (p : nat * list pattern) : option (nat * list pattern)
+ Definition filter_pattern_pident (pidc : pident) (p : nat * list pattern) : option (nat * list pattern)
:= match p with
- | (n, pattern.Wildcard _::ps)
- => Some (n, ps)
| (n, pattern.Ident pidc'::ps)
=> if pident_beq pidc pidc'
then Some (n, ps)
else None
+ | (_, pattern.Wildcard _::_)
| (_, pattern.App _ _::_)
| (_, nil)
=> None
@@ -628,13 +640,14 @@ Module Compilers.
=> (onfailure <- compile_rewrites ps;
Some (TryLeaf n1 onfailure))
| Some Datatypes.O
- => default_case <- compile_rewrites (filter_pattern_wildcard pattern_matrix);
+ => let '(pattern_matrix, default_pattern_matrix) := split_at_first_pattern_wildcard pattern_matrix in
+ default_case <- compile_rewrites default_pattern_matrix;
app_case <- (if contains_pattern_app pattern_matrix
- then option_map Some (compile_rewrites (Option.List.map refine_pattern_app pattern_matrix))
+ then option_map Some (compile_rewrites (Option.List.map filter_pattern_app pattern_matrix))
else Some None);
let pidcs := get_unique_pattern_ident pattern_matrix in
let icases := Option.List.map
- (fun pidc => option_map (pair pidc) (compile_rewrites (Option.List.map (refine_pattern_pident pidc) pattern_matrix)))
+ (fun pidc => option_map (pair pidc) (compile_rewrites (Option.List.map (filter_pattern_pident pidc) pattern_matrix)))
pidcs in
Some (Switch icases app_case default_case)
| Some i
@@ -901,8 +914,6 @@ Module Compilers.
:= fun var => @rewrite var (rewrite_head var) fuel t (e _).
End Compile.
- Module pident := pattern.ident.
-
Module Make.
Section make_rewrite_rules.
Import Compile.
@@ -932,8 +943,8 @@ Module Compilers.
Local Notation castv := (@castv ident var).
Definition make_base_Literal_pattern (t : base.type.base) : pattern
- := Eval cbv [pident.of_typed_ident] in
- pattern.Ident (pident.of_typed_ident (@ident.Literal t DefaultValue.type.base.default)).
+ := Eval cbv [pattern.ident.of_typed_ident] in
+ pattern.Ident (pattern.ident.of_typed_ident (@ident.Literal t DefaultValue.type.base.default)).
Definition bind_base_Literal_pattern (t : base.type.base) : binding_dataT (make_base_Literal_pattern t) ~> base.interp t
:= match t return binding_dataT (make_base_Literal_pattern t) ~> base.interp t with
@@ -952,7 +963,7 @@ Module Compilers.
b <- make_Literal_pattern B;
Some (existT
(fun p : pattern => binding_dataT p ~> base.interp (A * B))
- (#pident.pair @ (projT1 a) @ (projT1 b))%pattern
+ (#pattern.ident.pair @ (projT1 a) @ (projT1 b))%pattern
(fun '(args : unit * binding_dataT (projT1 a) * binding_dataT (projT1 b))
=> (av <--- projT2 a (snd (fst args));
bv <--- projT2 b (snd args);
@@ -983,8 +994,8 @@ Module Compilers.
Definition make_interp_rewrite'' {t} (idc : ident t) : option rewrite_ruleT
:= make_interp_rewrite'
t
- (pattern.Ident (pident.of_typed_ident idc))
- (fun iargs => return (ident.interp (pident.retype_ident idc iargs)))%cps.
+ (pattern.Ident (pattern.ident.of_typed_ident idc))
+ (fun iargs => return (ident.interp (pattern.ident.retype_ident idc iargs)))%cps.
(*
Definition make_interp_rewrite {t} (idc : ident t)
:= invert_Some (make_interp_rewrite'' idc).
@@ -1007,8 +1018,8 @@ Module Compilers.
| None => (eval cbv [List.rev List.app] in (List.rev so_far))
end.
Local Ltac make_valid_interp_rules :=
- let body := constr:(fun t idc => @pident.eta_ident_cps _ t idc (@make_interp_rewrite'')) in
- let body := (eval cbv [pident.eta_ident_cps make_interp_rewrite'' make_interp_rewrite' make_Literal_pattern pident.of_typed_ident Option.bind projT1 projT2 cpsbind cpsreturn cpscall ident.interp pident.retype_ident ident.gen_interp bind_base_Literal_pattern make_base_Literal_pattern] in body) in
+ let body := constr:(fun t idc => @pattern.ident.eta_ident_cps _ t idc (@make_interp_rewrite'')) in
+ let body := (eval cbv [pattern.ident.eta_ident_cps make_interp_rewrite'' make_interp_rewrite' make_Literal_pattern pattern.ident.of_typed_ident Option.bind projT1 projT2 cpsbind cpsreturn cpscall ident.interp pattern.ident.retype_ident ident.gen_interp bind_base_Literal_pattern make_base_Literal_pattern] in body) in
let body := (eval cbn [base.interp binding_dataT pattern.ident.arg_types base.base_interp ident.smart_Literal fold_right map] in body) in
let retv := get_all_valid_interp_rules_from body (@nil rewrite_ruleT) in
exact retv.
@@ -1081,10 +1092,10 @@ Module Compilers.
:= (let f' := (@lift_with_bindings p _ _ (fun x:@UnderLetsAnyExprCpsOpt base.type ident value => (x' <-- x; oret (existT (opt_anyexprP value) true x'))%cps) f%expr) in
make_rewrite'_cps p f').
- Local Notation "x' <- v ; C" := (fun T k => v%cps T (fun x' => match x' with Some x' => (C%cps : UnderLetsAnyExprCpsOpt) T k | None => k None end)) : cps_scope.
+ Local Notation "x' <- v ; C" := (fun T k => v%cps T (fun x' => Option.sequence_return (Compile.option_bind' x' (fun x' => Some ((C%cps : UnderLetsAnyExprCpsOpt) T k))) (k None))) : cps_scope.
Local Notation "x <-- y ; f" := (UnderLets.splice y (fun x => (f%cps : UnderLetsExpr _))) : cps_scope.
Local Notation "x <--- y ; f" := (UnderLets.splice_list y (fun x => (f%cps : UnderLetsExpr _))) : cps_scope.
- Local Notation "x <---- y ; f" := (fun T k => match y with Some x => (f%cps : UnderLetsAnyExprCpsOpt) T k | None => k None end) : cps_scope.
+ Local Notation "x <---- y ; f" := (fun T k => Option.sequence_return (Compile.option_bind' y (fun x => Some ((f%cps : UnderLetsAnyExprCpsOpt) T k))) (k None)) : cps_scope.
Definition rlist_rect {A P}
{ivar}
@@ -1157,46 +1168,46 @@ In the RHS, the follow notation applies:
:= Eval cbn [Make.interp_rewrite_rules List.app] in
Make.interp_rewrite_rules
++ [
- make_rewrite (#pident.fst @ (??, ??)) (fun _ x _ y => x)
- ; make_rewrite (#pident.snd @ (??, ??)) (fun _ x _ y => y)
- ; make_rewrite (#pident.List_repeat @ ?? @ #?ℕ) (fun _ x n => reify_list (repeat x n))
+ make_rewrite (#pattern.ident.fst @ (??, ??)) (fun _ x _ y => x)
+ ; make_rewrite (#pattern.ident.snd @ (??, ??)) (fun _ x _ y => y)
+ ; make_rewrite (#pattern.ident.List_repeat @ ?? @ #?ℕ) (fun _ x n => reify_list (repeat x n))
; make_rewrite
- (#pident.bool_rect @ ??{() -> ??} @ ??{() -> ??} @ #?𝔹)
+ (#pattern.ident.bool_rect @ ??{() -> ??} @ ??{() -> ??} @ #?𝔹)
(fun _ t _ f b
=> if b return UnderLetsExpr (type.base (if b then _ else _))
then t ##tt
else f ##tt)
; make_rewrite
- (#pident.prod_rect @ ??{?? -> ?? -> ??} @ (??, ??))
+ (#pattern.ident.prod_rect @ ??{?? -> ?? -> ??} @ (??, ??))
(fun _ _ _ f _ x _ y
=> x <- castbe x; y <- castbe y; ret (f x y))
; make_rewrite
(??{list ??} ++ ??{list ??})
(fun _ xs _ ys => rlist_rect_cast ys (fun x _ xs_ys => x :: xs_ys) xs)
; make_rewrite
- (#pident.List_firstn @ #?ℕ @ ??{list ??})
+ (#pattern.ident.List_firstn @ #?ℕ @ ??{list ??})
(fun n _ xs
=> xs <- reflect_list_cps xs;
reify_list (List.firstn n xs))
; make_rewrite
- (#pident.List_skipn @ #?ℕ @ ??{list ??})
+ (#pattern.ident.List_skipn @ #?ℕ @ ??{list ??})
(fun n _ xs
=> xs <- reflect_list_cps xs;
reify_list (List.skipn n xs))
; make_rewrite
- (#pident.List_rev @ ??{list ??})
+ (#pattern.ident.List_rev @ ??{list ??})
(fun _ xs
=> xs <- reflect_list_cps xs;
reify_list (List.rev xs))
; make_rewrite_step
- (#pident.List_flat_map @ ??{?? -> list ??} @ ??{list ??})
+ (#pattern.ident.List_flat_map @ ??{?? -> list ??} @ ??{list ??})
(fun _ B f _ xs
=> rlist_rect_cast
[]
(fun x _ flat_map_tl => fx <-- f x; UnderLets.Base ($fx ++ flat_map_tl))
xs)
; make_rewrite_step
- (#pident.List_partition @ ??{?? -> base.type.bool} @ ??{list ??})
+ (#pattern.ident.List_partition @ ??{?? -> base.type.bool} @ ??{list ??})
(fun _ f _ xs
=> rlist_rect_cast
([], [])
@@ -1210,7 +1221,7 @@ In the RHS, the follow notation applies:
@ partition_tl))
xs)
; make_rewrite
- (#pident.List_fold_right @ ??{?? -> ?? -> ??} @ ?? @ ??{list ??})
+ (#pattern.ident.List_fold_right @ ??{?? -> ?? -> ??} @ ?? @ ??{list ??})
(fun _ _ _ f B init A xs
=> f <- @castv _ (A -> B -> B)%etype f;
rlist_rect
@@ -1218,7 +1229,7 @@ In the RHS, the follow notation applies:
(fun x _ y => f x y)
xs)
; make_rewrite
- (#pident.list_rect @ ??{() -> ??} @ ??{?? -> ?? -> ?? -> ??} @ ??{list ??})
+ (#pattern.ident.list_rect @ ??{() -> ??} @ ??{?? -> ?? -> ?? -> ??} @ ??{list ??})
(fun P Pnil _ _ _ _ Pcons A xs
=> Pcons <- @castv _ (A -> base.type.list A -> P -> P) Pcons;
rlist_rect
@@ -1226,46 +1237,46 @@ In the RHS, the follow notation applies:
(fun x' xs' rec => Pcons x' (reify_list xs') rec)
xs)
; make_rewrite
- (#pident.list_case @ ??{() -> ??} @ ??{?? -> ?? -> ??} @ []) (fun _ Pnil _ _ _ Pcons => ret (Pnil ##tt))
+ (#pattern.ident.list_case @ ??{() -> ??} @ ??{?? -> ?? -> ??} @ []) (fun _ Pnil _ _ _ Pcons => ret (Pnil ##tt))
; make_rewrite
- (#pident.list_case @ ??{() -> ??} @ ??{?? -> ?? -> ??} @ (?? :: ??))
+ (#pattern.ident.list_case @ ??{() -> ??} @ ??{?? -> ?? -> ??} @ (?? :: ??))
(fun _ Pnil _ _ _ Pcons _ x _ xs
=> x <- castbe x; xs <- castbe xs; ret (Pcons x xs))
; make_rewrite
- (#pident.List_map @ ??{?? -> ??} @ ??{list ??})
+ (#pattern.ident.List_map @ ??{?? -> ??} @ ??{list ??})
(fun _ _ f _ xs
=> rlist_rect_cast
[]
(fun x _ fxs => fx <-- f x; fx :: fxs)
xs)
; make_rewrite
- (#pident.List_nth_default @ ?? @ ??{list ??} @ #?ℕ)
+ (#pattern.ident.List_nth_default @ ?? @ ??{list ??} @ #?ℕ)
(fun _ default _ ls n
=> default <- castbe default;
ls <- reflect_list_cps ls;
nth_default default ls n)
; make_rewrite
- (#pident.nat_rect @ ??{() -> ??} @ ??{base.type.nat -> ?? -> ??} @ #?ℕ)
+ (#pattern.ident.nat_rect @ ??{() -> ??} @ ??{base.type.nat -> ?? -> ??} @ #?ℕ)
(fun P O_case _ _ S_case n
=> S_case <- @castv _ (@type.base base.type base.type.nat -> type.base P -> type.base P) S_case;
ret (nat_rect _ (O_case ##tt) (fun n' rec => rec <-- rec; S_case ##n' rec) n))
; make_rewrite
- (#pident.nat_rect_arrow @ ??{?? -> ??} @ ??{base.type.nat -> (?? -> ??) -> (?? -> ??)} @ #?ℕ @ ??)
+ (#pattern.ident.nat_rect_arrow @ ??{?? -> ??} @ ??{base.type.nat -> (?? -> ??) -> (?? -> ??)} @ #?ℕ @ ??)
(fun P Q O_case _ _ _ _ S_case n _ v
=> S_case <- @castv _ (@type.base base.type base.type.nat -> (type.base P -> type.base Q) -> (type.base P -> type.base Q)) S_case;
v <- castbe v;
ret (nat_rect _ O_case (fun n' rec v => S_case ##n' rec v) n v))
; make_rewrite
- (#pident.List_length @ ??{list ??})
+ (#pattern.ident.List_length @ ??{list ??})
(fun _ xs => xs <- reflect_list_cps xs; ##(List.length xs))
; make_rewrite
- (#pident.List_combine @ ??{list ??} @ ??{list ??})
+ (#pattern.ident.List_combine @ ??{list ??} @ ??{list ??})
(fun _ xs _ ys
=> xs <- reflect_list_cps xs;
ys <- reflect_list_cps ys;
reify_list (List.map (fun '((x, y)%core) => (x, y)) (List.combine xs ys)))
; make_rewrite
- (#pident.List_update_nth @ #?ℕ @ ??{?? -> ??} @ ??{list ??})
+ (#pattern.ident.List_update_nth @ #?ℕ @ ??{?? -> ??} @ ??{list ??})
(fun n _ _ f A ls
=> f <- @castv _ (A -> A) f;
ls <- reflect_list_cps ls;
@@ -1278,8 +1289,8 @@ In the RHS, the follow notation applies:
]%list%pattern%cps%option%under_lets%Z%bool.
Definition arith_rewrite_rules : rewrite_rulesT
- := [make_rewrite (#pident.fst @ (??, ??)) (fun _ x _ y => x)
- ; make_rewrite (#pident.snd @ (??, ??)) (fun _ x _ y => y)
+ := [make_rewrite (#pattern.ident.fst @ (??, ??)) (fun _ x _ y => x)
+ ; make_rewrite (#pattern.ident.snd @ (??, ??)) (fun _ x _ y => y)
; make_rewrite (#?ℤ + ??{ℤ}) (fun z v => v when Z.eqb z 0)
; make_rewrite (??{ℤ} + #?ℤ ) (fun v z => v when Z.eqb z 0)
; make_rewrite (#?ℤ + (-??{ℤ})) (fun z v => ##z - v when Z.gtb z 0)
@@ -1328,31 +1339,31 @@ In the RHS, the follow notation applies:
; make_rewrite (??{ℤ} mod #?ℤ) (fun x y => x &' ##(y-1)%Z when (y =? (2^Z.log2 y)))
; make_rewrite (-(-??{ℤ})) (fun v => v)
- ; 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)
- ; make_rewrite (#pident.Z_mul_split @ #?ℤ @ ??{ℤ} @ #?ℤ) (fun s y xx => (y, ##0)%Z when Z.eqb xx 1)
- ; 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 (#pattern.ident.Z_mul_split @ #?ℤ @ #?ℤ @ ??{ℤ}) (fun s xx y => (##0, ##0)%Z when Z.eqb xx 0)
+ ; make_rewrite (#pattern.ident.Z_mul_split @ #?ℤ @ ??{ℤ} @ #?ℤ) (fun s y xx => (##0, ##0)%Z when Z.eqb xx 0)
+ ; make_rewrite (#pattern.ident.Z_mul_split @ #?ℤ @ #?ℤ @ ??{ℤ}) (fun s xx y => (y, ##0)%Z when Z.eqb xx 1)
+ ; make_rewrite (#pattern.ident.Z_mul_split @ #?ℤ @ ??{ℤ} @ #?ℤ) (fun s y xx => (y, ##0)%Z when Z.eqb xx 1)
+ ; make_rewrite (#pattern.ident.Z_mul_split @ #?ℤ @ #?ℤ @ ??{ℤ}) (fun s xx y => (-y, ##0%Z) when Z.eqb xx (-1))
+ ; make_rewrite (#pattern.ident.Z_mul_split @ #?ℤ @ ??{ℤ} @ #?ℤ) (fun s y xx => (-y, ##0%Z) when Z.eqb xx (-1))
; make_rewrite
- (#pident.Z_add_get_carry @ ??{ℤ} @ (-??{ℤ}) @ ??{ℤ})
+ (#pattern.ident.Z_add_get_carry @ ??{ℤ} @ (-??{ℤ}) @ ??{ℤ})
(fun s y x => ret (UnderLets.UnderLet
(#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 @ ??{ℤ} @ ??{ℤ} @ (-??{ℤ}))
+ (#pattern.ident.Z_add_get_carry @ ??{ℤ} @ ??{ℤ} @ (-??{ℤ}))
(fun s x y => ret (UnderLets.UnderLet
(#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 @ ??{ℤ} @ #?ℤ @ ??{ℤ})
+ (#pattern.ident.Z_add_get_carry @ ??{ℤ} @ #?ℤ @ ??{ℤ})
(fun s yy x => ret (UnderLets.UnderLet
(#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 @ ??{ℤ} @ ??{ℤ} @ #?ℤ)
+ (#pattern.ident.Z_add_get_carry @ ??{ℤ} @ ??{ℤ} @ #?ℤ)
(fun s x yy => ret (UnderLets.UnderLet
(#ident.Z_sub_get_borrow @ s @ x @ ##(-yy)%Z)
(fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))
@@ -1360,85 +1371,85 @@ In the RHS, the follow notation applies:
; make_rewrite
- (#pident.Z_add_with_get_carry @ ??{ℤ} @ (-??{ℤ}) @ (-??{ℤ}) @ ??{ℤ})
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ (-??{ℤ}) @ (-??{ℤ}) @ ??{ℤ})
(fun s c y x => ret (UnderLets.UnderLet
(#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 @ ??{ℤ} @ (-??{ℤ}) @ ??{ℤ} @ (-??{ℤ}))
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ (-??{ℤ}) @ ??{ℤ} @ (-??{ℤ}))
(fun s c x y => ret (UnderLets.UnderLet
(#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 @ ??{ℤ} @ #?ℤ @ (-??{ℤ}) @ ??{ℤ})
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ #?ℤ @ (-??{ℤ}) @ ??{ℤ})
(fun s cc y x => ret (UnderLets.UnderLet
(#ident.Z_sub_get_borrow @ s @ x @ y)
(fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))
when cc =? 0)
; make_rewrite
- (#pident.Z_add_with_get_carry @ ??{ℤ} @ #?ℤ @ (-??{ℤ}) @ ??{ℤ})
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ #?ℤ @ (-??{ℤ}) @ ??{ℤ})
(fun s cc y x => ret (UnderLets.UnderLet
(#ident.Z_sub_with_get_borrow @ s @ ##(-cc)%Z @ x @ y)
(fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))
when cc <? 0)
; make_rewrite
- (#pident.Z_add_with_get_carry @ ??{ℤ} @ #?ℤ @ ??{ℤ} @ (-??{ℤ}))
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ #?ℤ @ ??{ℤ} @ (-??{ℤ}))
(fun s cc x y => ret (UnderLets.UnderLet
(#ident.Z_sub_get_borrow @ s @ x @ y)
(fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))
when cc =? 0)
; make_rewrite
- (#pident.Z_add_with_get_carry @ ??{ℤ} @ #?ℤ @ ??{ℤ} @ (-??{ℤ}))
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ #?ℤ @ ??{ℤ} @ (-??{ℤ}))
(fun s cc x y => ret (UnderLets.UnderLet
(#ident.Z_sub_with_get_borrow @ s @ ##(-cc)%Z @ x @ y)
(fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))
when cc <? 0)
; make_rewrite
- (#pident.Z_add_with_get_carry @ ??{ℤ} @ (-??{ℤ}) @ #?ℤ @ ??{ℤ})
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ (-??{ℤ}) @ #?ℤ @ ??{ℤ})
(fun s c yy x => ret (UnderLets.UnderLet
(#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 @ ??{ℤ} @ (-??{ℤ}) @ ??{ℤ} @ #?ℤ)
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ (-??{ℤ}) @ ??{ℤ} @ #?ℤ)
(fun s c x yy => ret (UnderLets.UnderLet
(#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 @ ??{ℤ} @ #?ℤ @ #?ℤ @ ??{ℤ})
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ #?ℤ @ #?ℤ @ ??{ℤ})
(fun s cc yy x => ret (UnderLets.UnderLet
(#ident.Z_sub_with_get_borrow @ s @ ##(-cc)%Z @ x @ ##(-yy)%Z)
(fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))
when (yy <=? 0) && (cc <=? 0) && ((yy + cc) <? 0)) (* at least one must be strictly negative *)
; make_rewrite
- (#pident.Z_add_with_get_carry @ ??{ℤ} @ #?ℤ @ ??{ℤ} @ #?ℤ)
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ #?ℤ @ ??{ℤ} @ #?ℤ)
(fun s cc x yy => ret (UnderLets.UnderLet
(#ident.Z_sub_with_get_borrow @ s @ ##(-cc)%Z @ x @ ##(-yy)%Z)
(fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))
when (yy <=? 0) && (cc <=? 0) && ((yy + cc) <? 0)) (* at least one must be strictly negative *)
- ; 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 (#pattern.ident.Z_add_get_carry @ ??{ℤ} @ #?ℤ @ ??{ℤ}) (fun s xx y => (y, ##0) when xx =? 0)
+ ; make_rewrite (#pattern.ident.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_step (#pident.Z_add_with_carry @ ??{ℤ} @ ??{ℤ} @ ??{ℤ}) (fun x y z => $x + $y + $z)*)
+ ; make_rewrite (#pattern.ident.Z_add_with_carry @ #?ℤ @ ??{ℤ} @ ??{ℤ}) (fun cc x y => x + y when cc =? 0)
+ (*; make_rewrite_step (#pattern.ident.Z_add_with_carry @ ??{ℤ} @ ??{ℤ} @ ??{ℤ}) (fun x y z => $x + $y + $z)*)
; make_rewrite
- (#pident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ #?ℤ @ ??{ℤ}) (fun s cc xx y => (y, ##0) when (cc =? 0) && (xx =? 0))
+ (#pattern.ident.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))
+ (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ ??{ℤ} @ #?ℤ) (fun s cc y xx => (y, ##0) when (cc =? 0) && (xx =? 0))
(*; make_rewrite
- (#pident.Z_add_with_get_carry @ ??{ℤ} @ ??{ℤ} @ #?ℤ @ #?ℤ) (fun s c xx yy => (c, ##0) when (xx =? 0) && (yy =? 0))*)
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ ??{ℤ} @ #?ℤ @ #?ℤ) (fun s c xx yy => (c, ##0) when (xx =? 0) && (yy =? 0))*)
; make_rewrite (* carry = 0: ADC x y -> ADD x y *)
- (#pident.Z_add_with_get_carry @ ??{ℤ} @ #?ℤ @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ #?ℤ @ ??{ℤ} @ ??{ℤ})
(fun s cc x y => ret (UnderLets.UnderLet
(#ident.Z_add_get_carry @ s @ x @ y)
(fun vc => UnderLets.Base (#ident.fst @ $vc, #ident.snd @ $vc)))
when cc =? 0)
; make_rewrite (* ADC 0 0 -> (ADX 0 0, 0) *) (* except we don't do ADX, because C stringification doesn't handle it *)
- (#pident.Z_add_with_get_carry @ ??{ℤ} @ ??{ℤ} @ #?ℤ @ #?ℤ)
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ ??{ℤ} @ #?ℤ @ #?ℤ)
(fun s c xx yy => ret (UnderLets.UnderLet
(#ident.Z_add_with_get_carry @ s @ c @ ##xx @ ##yy)
(fun vc => UnderLets.Base (#ident.fst @ $vc, ##0)))
@@ -1447,33 +1458,33 @@ In the RHS, the follow notation applies:
(* let-bind any adc/sbb/mulx *)
; make_rewrite
- (#pident.Z_add_with_get_carry @ ??{ℤ} @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.Z_add_with_get_carry @ ??{ℤ} @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
(fun s c x y => ret (UnderLets.UnderLet (#ident.Z_add_with_get_carry @ s @ c @ x @ y)
(fun vc => UnderLets.Base (#ident.fst @ $vc, #ident.snd @ $vc))))
; make_rewrite
- (#pident.Z_add_with_carry @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.Z_add_with_carry @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
(fun c x y => ret (UnderLets.UnderLet (#ident.Z_add_with_carry @ c @ x @ y)
(fun vc => UnderLets.Base ($vc))))
; make_rewrite
- (#pident.Z_add_get_carry @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.Z_add_get_carry @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
(fun s x y => ret (UnderLets.UnderLet (#ident.Z_add_get_carry @ s @ x @ y)
(fun vc => UnderLets.Base (#ident.fst @ $vc, #ident.snd @ $vc))))
; make_rewrite
- (#pident.Z_sub_with_get_borrow @ ??{ℤ} @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.Z_sub_with_get_borrow @ ??{ℤ} @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
(fun s c x y => ret (UnderLets.UnderLet (#ident.Z_sub_with_get_borrow @ s @ c @ x @ y)
(fun vc => UnderLets.Base (#ident.fst @ $vc, #ident.snd @ $vc))))
; make_rewrite
- (#pident.Z_sub_get_borrow @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.Z_sub_get_borrow @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
(fun s x y => ret (UnderLets.UnderLet (#ident.Z_sub_get_borrow @ s @ x @ y)
(fun vc => UnderLets.Base (#ident.fst @ $vc, #ident.snd @ $vc))))
; make_rewrite
- (#pident.Z_mul_split @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.Z_mul_split @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
(fun s x y => ret (UnderLets.UnderLet (#ident.Z_mul_split @ s @ x @ y)
(fun v => UnderLets.Base (#ident.fst @ $v, #ident.snd @ $v))))
; 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))
+ (#pattern.ident.Z_cast2 @ (??{ℤ}, ??{ℤ})) (fun r x y => (#(ident.Z_cast (fst r)) @ $x, #(ident.Z_cast (snd r)) @ $y))
; make_rewrite (-??{ℤ}) (fun e => ret (UnderLets.UnderLet e (fun v => UnderLets.Base (-$v))) when negb (SubstVarLike.is_var_fst_snd_pair_opp e)) (* inline negation when the rewriter wouldn't already inline it *)
]%list%pattern%cps%option%under_lets%Z%bool.
@@ -1530,19 +1541,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 @ #?ℤ @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ} @ #?ℤ))
+ (#pattern.ident.Z_add_get_carry @ #?ℤ @ ??{ℤ} @ (#pattern.ident.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 @ #?ℤ @ (#pident.Z_shiftl @ ??{ℤ} @ #?ℤ) @ ??{ℤ})
+ (#pattern.ident.Z_add_get_carry @ #?ℤ @ (#pattern.ident.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 @ #?ℤ @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ))
+ (#pattern.ident.Z_add_get_carry @ #?ℤ @ ??{ℤ} @ (#pattern.ident.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 @ #?ℤ @ (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ) @ ??{ℤ})
+ (#pattern.ident.Z_add_get_carry @ #?ℤ @ (#pattern.ident.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 @ #?ℤ @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.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)
@@ -1552,19 +1563,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 @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ} @ #?ℤ))
+ (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ (#pattern.ident.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 @ #?ℤ @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ} @ #?ℤ) @ ??{ℤ})
+ (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??{ℤ} @ (#pattern.ident.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 @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ))
+ (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ (#pattern.ident.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 @ #?ℤ @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ) @ ??{ℤ})
+ (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??{ℤ} @ (#pattern.ident.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 @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.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)
@@ -1572,13 +1583,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 @ #?ℤ @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ} @ #?ℤ))
+ (#pattern.ident.Z_sub_get_borrow @ #?ℤ @ ??{ℤ} @ (#pattern.ident.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 @ #?ℤ @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ))
+ (#pattern.ident.Z_sub_get_borrow @ #?ℤ @ ??{ℤ} @ (#pattern.ident.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 @ #?ℤ @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.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)
@@ -1586,17 +1597,17 @@ 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 @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ} @ #?ℤ))
+ (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ (#pattern.ident.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 @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ))
+ (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ (#pattern.ident.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 @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.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 @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ #?ℤ)
+ (#pattern.ident.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)
@@ -1604,20 +1615,20 @@ 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 @ #?ℤ @ ??{ℤ}) @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.Z_zselect @ (#pattern.ident.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 @ #?ℤ @ ??{ℤ}) @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.Z_zselect @ (#pattern.ident.Z_land @ #?ℤ @ ??{ℤ}) @ ??{ℤ} @ ??{ℤ})
(fun mask c x y => #ident.fancy_sell @ (c, x, y) when mask =? 1)
; make_rewrite
- (#pident.Z_zselect @ (#pident.Z_land @ ??{ℤ} @ #?ℤ) @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.Z_zselect @ (#pattern.ident.Z_land @ ??{ℤ} @ #?ℤ) @ ??{ℤ} @ ??{ℤ})
(fun c mask x y => #ident.fancy_sell @ (c, x, y) when mask =? 1)
; make_rewrite
- (#pident.Z_zselect @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.Z_zselect @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
(fun c x y => #ident.fancy_selc @ (c, x, y))
(*Z.add_modulo @@ (?x, ?y, ?m) --> addm @@ (x, y, m)*)
; make_rewrite
- (#pident.Z_add_modulo @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
+ (#pattern.ident.Z_add_modulo @ ??{ℤ} @ ??{ℤ} @ ??{ℤ})
(fun x y m => #ident.fancy_addm @ (x, y, m))
(*
Z.mul @@ (?x &' (2^128-1), ?y &' (2^128-1)) --> mulll @@ (x, y)
@@ -1627,71 +1638,71 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
*)
(* literal on left *)
; make_rewrite
- (#?ℤ * (#pident.Z_land @ ??{ℤ} @ #?ℤ))
+ (#?ℤ * (#pattern.ident.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 @ #?ℤ @ ??{ℤ}))
+ (#?ℤ * (#pattern.ident.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 @ ??{ℤ} @ #?ℤ))
+ (#?ℤ * (#pattern.ident.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 @ #?ℤ @ ??{ℤ}))
+ (#?ℤ * (#pattern.ident.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_land @ ??{ℤ} @ #?ℤ))
+ (#?ℤ * (#pattern.ident.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 @ ??{ℤ} @ #?ℤ))
+ (#?ℤ * (#pattern.ident.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 @ #?ℤ @ ??{ℤ}) * #?ℤ)
+ ((#pattern.ident.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 @ ??{ℤ} @ #?ℤ) * #?ℤ)
+ ((#pattern.ident.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 @ #?ℤ @ ??{ℤ}) * #?ℤ)
+ ((#pattern.ident.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_land @ ??{ℤ} @ #?ℤ) * #?ℤ)
+ ((#pattern.ident.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 @ ??{ℤ} @ #?ℤ) * #?ℤ)
+ ((#pattern.ident.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 @ ??{ℤ} @ #?ℤ) * #?ℤ)
+ ((#pattern.ident.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 @ #?ℤ @ ??{ℤ}))
+ ((#pattern.ident.Z_land @ #?ℤ @ ??{ℤ}) * (#pattern.ident.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_land @ #?ℤ @ ??{ℤ}))
+ ((#pattern.ident.Z_land @ ??{ℤ} @ #?ℤ) * (#pattern.ident.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 @ ??{ℤ} @ #?ℤ))
+ ((#pattern.ident.Z_land @ #?ℤ @ ??{ℤ}) * (#pattern.ident.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 @ ??{ℤ} @ #?ℤ))
+ ((#pattern.ident.Z_land @ ??{ℤ} @ #?ℤ) * (#pattern.ident.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 @ ??{ℤ} @ #?ℤ))
+ ((#pattern.ident.Z_land @ #?ℤ @ ??{ℤ}) * (#pattern.ident.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 @ ??{ℤ} @ #?ℤ))
+ ((#pattern.ident.Z_land @ ??{ℤ} @ #?ℤ) * (#pattern.ident.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 @ #?ℤ @ ??{ℤ}))
+ ((#pattern.ident.Z_shiftr @ ??{ℤ} @ #?ℤ) * (#pattern.ident.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 @ ??{ℤ} @ #?ℤ))
+ ((#pattern.ident.Z_shiftr @ ??{ℤ} @ #?ℤ) * (#pattern.ident.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 @ ??{ℤ} @ #?ℤ))
+ ((#pattern.ident.Z_shiftr @ ??{ℤ} @ #?ℤ) * (#pattern.ident.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.
@@ -1724,7 +1735,10 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
:= Eval cbv -[fancy_pr2_rewrite_rules
base.interp base.try_make_transport_cps
type.try_make_transport_cps type.try_transport_cps
+ pattern.ident.invert_bind_args
+ Let_In Option.sequence Option.sequence_return
UnderLets.splice UnderLets.to_expr
+ Compile.option_bind'
Compile.reflect Compile.reify Compile.reify_and_let_binds_cps UnderLets.reify_and_let_binds_base_cps
Compile.value' SubstVarLike.is_var_fst_snd_pair_opp
] in @fancy_rewrite_head0 var invert_low invert_high do_again t idc.
@@ -1749,6 +1763,7 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
Compile.lift_pbase_type_interp_cps
Compile.lift_ptype_interp_cps
Compile.lift_with_bindings
+ Compile.option_bind'
Compile.pbase_type_interp_cps
Compile.ptype_interp
Compile.ptype_interp_cps
@@ -1772,13 +1787,20 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
Local Arguments base.try_make_base_transport_cps _ !_ !_.
Local Arguments base.try_make_transport_cps _ !_ !_.
Local Arguments type.try_make_transport_cps _ _ _ !_ !_.
+ Local Arguments Option.sequence {A} !v1 v2.
+ Local Arguments Option.sequence_return {A} !v1 v2.
+ Local Arguments Option.bind {A B} !_ _.
+ Local Arguments pattern.ident.invert_bind_args {t} !_ !_.
+ Local Arguments base.try_make_transport_cps {P} t1 t2 {_} _.
Local Arguments fancy_rewrite_head2 / .
Time Definition fancy_rewrite_head
:= Eval cbn [id
fancy_rewrite_head2
cpsbind cpscall cps_option_bind cpsreturn
+ pattern.ident.invert_bind_args
Compile.reify Compile.reify_and_let_binds_cps Compile.reflect Compile.value'
+ Option.sequence Option.sequence_return Option.bind
UnderLets.reify_and_let_binds_base_cps
UnderLets.splice UnderLets.splice_list UnderLets.to_expr
base.interp base.base_interp
@@ -1787,6 +1809,9 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
] in fancy_rewrite_head2.
(* Finished transaction in 13.298 secs (13.283u,0.s) (successful) *)
+ Local Set Printing Depth 1000000.
+ Local Set Printing Width 200.
+ Local Notation "'llet' x := v 'in' f" := (Let_In v (fun x => f)).
Redirect "/tmp/fancy_rewrite_head" Print fancy_rewrite_head.
End red_fancy.
@@ -1800,7 +1825,10 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
:= Eval cbv -[nbe_pr2_rewrite_rules
base.interp base.try_make_transport_cps
type.try_make_transport_cps type.try_transport_cps
+ pattern.ident.invert_bind_args
+ Let_In Option.sequence Option.sequence_return
UnderLets.splice UnderLets.to_expr
+ Compile.option_bind'
Compile.reflect UnderLets.reify_and_let_binds_base_cps Compile.reify Compile.reify_and_let_binds_cps
Compile.value'
SubstVarLike.is_var_fst_snd_pair_opp
@@ -1823,6 +1851,7 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
Compile.eval_decision_tree
Compile.eval_rewrite_rules
Compile.expr_of_rawexpr
+ Compile.option_bind'
Compile.lift_pbase_type_interp_cps
Compile.lift_ptype_interp_cps
Compile.lift_with_bindings
@@ -1849,13 +1878,20 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
Local Arguments base.try_make_base_transport_cps _ !_ !_.
Local Arguments base.try_make_transport_cps _ !_ !_.
Local Arguments type.try_make_transport_cps _ _ _ !_ !_.
+ Local Arguments Option.sequence {A} !v1 v2.
+ Local Arguments Option.sequence_return {A} !v1 v2.
+ Local Arguments Option.bind {A B} !_ _.
+ Local Arguments pattern.ident.invert_bind_args {t} !_ !_.
+ Local Arguments base.try_make_transport_cps {P} t1 t2 {_} _.
Local Arguments nbe_rewrite_head2 / .
Time Definition nbe_rewrite_head
:= Eval cbn [id
nbe_rewrite_head2
cpsbind cpscall cps_option_bind cpsreturn
+ pattern.ident.invert_bind_args
Compile.reify Compile.reify_and_let_binds_cps Compile.reflect Compile.value'
+ Option.sequence Option.sequence_return Option.bind
UnderLets.reify_and_let_binds_base_cps
UnderLets.splice UnderLets.splice_list UnderLets.to_expr
base.interp base.base_interp
@@ -1864,6 +1900,9 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
] in nbe_rewrite_head2.
(* Finished transaction in 16.561 secs (16.54u,0.s) (successful) *)
+ Local Set Printing Depth 1000000.
+ Local Set Printing Width 200.
+ Local Notation "'llet' x := v 'in' f" := (Let_In v (fun x => f)).
Redirect "/tmp/nbe_rewrite_head" Print nbe_rewrite_head.
End red_nbe.
@@ -1877,7 +1916,10 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
:= Eval cbv -[arith_pr2_rewrite_rules
base.interp base.try_make_transport_cps
type.try_make_transport_cps type.try_transport_cps
+ pattern.ident.invert_bind_args
+ Let_In Option.sequence Option.sequence_return
UnderLets.splice UnderLets.to_expr
+ Compile.option_bind'
Compile.reflect UnderLets.reify_and_let_binds_base_cps Compile.reify Compile.reify_and_let_binds_cps
Compile.value'
SubstVarLike.is_var_fst_snd_pair_opp
@@ -1900,6 +1942,7 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
Compile.eval_decision_tree
Compile.eval_rewrite_rules
Compile.expr_of_rawexpr
+ Compile.option_bind'
Compile.lift_pbase_type_interp_cps
Compile.lift_ptype_interp_cps
Compile.lift_with_bindings
@@ -1926,13 +1969,20 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
Local Arguments base.try_make_base_transport_cps _ !_ !_.
Local Arguments base.try_make_transport_cps _ !_ !_.
Local Arguments type.try_make_transport_cps _ _ _ !_ !_.
+ Local Arguments Option.sequence {A} !v1 v2.
+ Local Arguments Option.sequence_return {A} !v1 v2.
+ Local Arguments Option.bind {A B} !_ _.
+ Local Arguments pattern.ident.invert_bind_args {t} !_ !_.
+ Local Arguments base.try_make_transport_cps {P} t1 t2 {_} _.
Local Arguments arith_rewrite_head2 / .
Time Definition arith_rewrite_head
:= Eval cbn [id
arith_rewrite_head2
cpsbind cpscall cps_option_bind cpsreturn
+ pattern.ident.invert_bind_args
Compile.reify Compile.reify_and_let_binds_cps Compile.reflect Compile.value'
+ Option.sequence Option.sequence_return Option.bind
UnderLets.reify_and_let_binds_base_cps
UnderLets.splice UnderLets.splice_list UnderLets.to_expr
base.interp base.base_interp
@@ -1941,6 +1991,9 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
] in arith_rewrite_head2.
(* Finished transaction in 16.561 secs (16.54u,0.s) (successful) *)
+ Local Set Printing Depth 1000000.
+ Local Set Printing Width 200.
+ Local Notation "'llet' x := v 'in' f" := (Let_In v (fun x => f)).
Redirect "/tmp/arith_rewrite_head" Print arith_rewrite_head.
End red_arith.
diff --git a/src/Experiments/NewPipeline/StandaloneHaskellMain.v b/src/Experiments/NewPipeline/StandaloneHaskellMain.v
index 131907bc1..7a8e5cc74 100644
--- a/src/Experiments/NewPipeline/StandaloneHaskellMain.v
+++ b/src/Experiments/NewPipeline/StandaloneHaskellMain.v
@@ -34,6 +34,8 @@ Extract Inlined Constant _IO_bind => "(Prelude.>>=)".
Extract Inlined Constant _IO_return => "return".
Extract Inlined Constant IO_unit => "GHC.Base.IO ()".
Extract Inlined Constant cast_io => "".
+(** Error messages from deeply nested axioms *)
+Extract Constant Rewriter.Compilers.RewriteRules.Compile.ERROR_BAD_REWRITE_RULE => "\pat e -> Prelude.error ""ERROR BAD REWRITE RULE""".
Local Notation "x <- y ; f" := (_IO_bind _ _ y (fun x => f)).
diff --git a/src/Experiments/NewPipeline/StandaloneOCamlMain.v b/src/Experiments/NewPipeline/StandaloneOCamlMain.v
index c5b4db570..0e41cdaa2 100644
--- a/src/Experiments/NewPipeline/StandaloneOCamlMain.v
+++ b/src/Experiments/NewPipeline/StandaloneOCamlMain.v
@@ -40,6 +40,8 @@ Extract Inlined Constant string_get => "String.get".
Extract Constant sys_argv => "Array.to_list Sys.argv".
Extract Inlined Constant string_init => "String.init".
Extract Constant raise_Failure => "fun x -> raise (Failure x)".
+(** Error messages from deeply nested axioms *)
+Extract Constant Rewriter.Compilers.RewriteRules.Compile.ERROR_BAD_REWRITE_RULE => "fun pat e -> failwith ""ERROR BAD REWRITE RULE""".
Fixpoint nat_of_int (x : int) : nat
:= match x with
diff --git a/src/Experiments/NewPipeline/arith_rewrite_head.out b/src/Experiments/NewPipeline/arith_rewrite_head.out
index 23b5205d2..e71a26c50 100644
--- a/src/Experiments/NewPipeline/arith_rewrite_head.out
+++ b/src/Experiments/NewPipeline/arith_rewrite_head.out
@@ -1,5 +1,5 @@
arith_rewrite_head =
-match idc in (ident t) return (Compile.value' true t) with
+match idc in (ident t) return (value' true t) with
| @ident.Literal t v =>
match
t as t0
@@ -8,10 +8,10 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base t0)))
with
- | base.type.unit => fun v0 : unit => UnderLets.Base ##(v0)%expr
- | base.type.Z => fun v0 : Z => UnderLets.Base ##(v0)%expr
- | base.type.bool => fun v0 : bool => UnderLets.Base ##(v0)%expr
- | base.type.nat => fun v0 : nat => UnderLets.Base ##(v0)%expr
+ | base.type.unit => fun v0 : unit => UnderLets.Base (##v0)%expr
+ | base.type.Z => fun v0 : Z => UnderLets.Base (##v0)%expr
+ | base.type.bool => fun v0 : bool => UnderLets.Base (##v0)%expr
+ | base.type.nat => fun v0 : nat => UnderLets.Base (##v0)%expr
end v
| ident.Nat_succ =>
fun x : defaults.expr (type.base base.type.nat) =>
@@ -41,132 +41,120 @@ match idc in (ident t) return (Compile.value' true t) with
=> UnderLets.Base (x, x0)%expr_pat
| @ident.fst A B =>
fun x : defaults.expr (type.base (A * B)%etype) =>
- match x with
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x1) x0 =>
- match
- match idc with
- | @ident.pair A0 B0 => Some (A0, B0)
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base A)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base A)))
- with
- | type.base t3 =>
- fun _ : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- A
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base A)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base A)) =>
- match a with
- | Some x' => UnderLets.Base (x' v)
- | None =>
- UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1
- => UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- end (Compile.reflect x0)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1 =>
- UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- end (Compile.reflect x1)
- | None => UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- end
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App _
- _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ | @expr.App
- _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ | @expr.App _ _ _
- s _ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ =>
- UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr _ |
- @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- | _ => UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- end
+ ((match x with
+ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc) x1) x0 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.pair;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base A))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ match
+ s as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base A))))
+ with
+ | type.base t3 =>
+ fun _ : defaults.expr (type.base t3) =>
+ Some
+ (base.try_make_transport_cps t2 A
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base A)) =>
+ match a with
+ | Some x' => UnderLets.Base (x' v)
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.fst) @ (??, ??))
+ (#(ident.fst)%expr @ x)%expr_pat)
+ end))
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x0)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
+ _ _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ |
+ @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ |
+ @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ => None
+ | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr
+ _ | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s
+ _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat)%option
| @ident.snd A B =>
fun x : defaults.expr (type.base (A * B)%etype) =>
- match x with
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x1) x0 =>
- match
- match idc with
- | @ident.pair A0 B0 => Some (A0, B0)
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base B)))
- with
- | type.base t2 =>
- fun _ : defaults.expr (type.base t2) =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base B)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t3
- B
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base B)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr (type.base B)) =>
- match a with
- | Some x' => UnderLets.Base (x' v0)
- | None =>
- UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1
- => UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- end (Compile.reflect x0)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1 =>
- UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- end (Compile.reflect x1)
- | None => UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- end
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App _
- _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ | @expr.App
- _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ | @expr.App _ _ _
- s _ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ =>
- UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr _ |
- @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- | _ => UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- end
+ ((match x with
+ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc) x1) x0 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.pair;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base B))))
+ with
+ | type.base t2 =>
+ fun _ : defaults.expr (type.base t2) =>
+ match
+ s as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base B))))
+ with
+ | type.base t3 =>
+ fun v0 : defaults.expr (type.base t3) =>
+ Some
+ (base.try_make_transport_cps t3 B
+ (fun
+ a : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr (type.base B)) =>
+ match a with
+ | Some x' => UnderLets.Base (x' v0)
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.snd) @ (??, ??))
+ (#(ident.snd)%expr @ x)%expr_pat)
+ end))
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x0)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
+ _ _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ |
+ @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ |
+ @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ => None
+ | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr
+ _ | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s
+ _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat)%option
| @ident.prod_rect A B T =>
fun
(x : defaults.expr (type.base A) ->
@@ -354,3786 +342,900 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base (#(ident.List_nth_default)%expr @ x @ x0 @ x1)%expr_pat
| ident.Z_add =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base x0
- else
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if args0 =? 0
- then UnderLets.Base x
- else UnderLets.Base (x + x0)%expr
- | None => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s _ #(idc0)%expr_pat x1 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args >? 0
- then UnderLets.Base (##(args) - x' v)%expr
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.Base
- (- (##((- args)%Z) + x'0 v0))%expr
- else
- match
- s as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v1 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr
- (type.base t1)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a1 with
- | Some x'1 =>
- UnderLets.Base
- (x - x'1 v1)%expr
- | None =>
- UnderLets.Base
- (x + x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false
- s0 ->
- Compile.value' true d0
- =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None =>
- UnderLets.Base (x + x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => UnderLets.Base (x + x0)%expr
- | _ => UnderLets.Base (x + x0)%expr
- end
- | None =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base x
- else UnderLets.Base (x + x0)%expr
- | None => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s _ #(idc0)%expr_pat x1 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x - x' v)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => UnderLets.Base (x + x0)%expr
- | _ => UnderLets.Base (x + x0)%expr
- end
- end
- | ($_)%expr =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t1 v0 =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base x
- else UnderLets.Base (x + x0)%expr
- | None => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s _ #(idc)%expr_pat x1 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1)) t2
- base.type.Z
+ (((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 0 then Some (UnderLets.Base x0) else None)
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 0 then Some (UnderLets.Base x) else None);;
+ match x with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t0 idc0) x1 =>
+ _ <- pattern.ident.invert_bind_args idc0 pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
(UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x - x' v0)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun _ : Compile.value' false s0 -> Compile.value' true d0
- => UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat
- _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x + x0)%expr
- | _ => UnderLets.Base (x + x0)%expr
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base x
- else UnderLets.Base (x + x0)%expr
- | None => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s0 _ #(idc)%expr_pat x1 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x - x' v)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1
- => UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0 _ (_ @ _)%expr_pat
- _ | @expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x + x0)%expr
- | _ => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s _ f x1 =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base x
- else
- match f with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args >? 0
- then
- UnderLets.Base (##(args) - x' v)%expr
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.Base
- (-
- (x'0 v0 + ##((- args)%Z)))%expr
- else
- match
- s as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v1 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr
- (type.base t1)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a1 with
- | Some x'1 =>
- UnderLets.Base
- (x0 - x'1 v1)%expr
- | None =>
- UnderLets.Base
- (x + x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value'
- false s0 ->
- Compile.value'
- true d0 =>
- UnderLets.Base
- (x + x0)%expr
- end (Compile.reflect x1)
- | None =>
- UnderLets.Base (x + x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end
- | _ => UnderLets.Base (x + x0)%expr
- end
- | None =>
- match f with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' => UnderLets.Base (x0 - x' v)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end
- | _ => UnderLets.Base (x + x0)%expr
- end
- end
- | ($_)%expr =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x0 - x' v0)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end
- | _ => UnderLets.Base (x + x0)%expr
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x0 - x' v)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end
- | _ => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s0 _ #(idc)%expr_pat x2 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match f with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- UnderLets.Base
- (- (x' v + x'0 v0))%expr
- | None =>
- UnderLets.Base (x + x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' => UnderLets.Base (x - x' v)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x2)
- end
- | ($_)%expr =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x - x' v0)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x2)
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x - x' v)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x2)
- | _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x - x' v)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x2)
- end
- | None =>
- match f with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' => UnderLets.Base (x0 - x' v)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end
- | _ => UnderLets.Base (x + x0)%expr
- end
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x0 - x' v0)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end
- | _ => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0
- _ (_ @ _)%expr_pat _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x0 - x' v)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end
- | _ => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x0 - x' v)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end
- | _ => UnderLets.Base (x + x0)%expr
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x0 - x' v)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x + x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x + x0)%expr
- end
- | _ => UnderLets.Base (x + x0)%expr
- end
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base x
- else UnderLets.Base (x + x0)%expr
- | None => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s _ #(idc)%expr_pat x2 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args >? 0
+ then Some (UnderLets.Base (##args - x' v)%expr)
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1);;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
(UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x - x' v)%expr
- | None => UnderLets.Base (x + x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun _ : Compile.value' false s0 -> Compile.value' true d0
- => UnderLets.Base (x + x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x + x0)%expr
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat
- _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x + x0)%expr
- | _ => UnderLets.Base (x + x0)%expr
- end
- end
-| ident.Z_mul =>
- fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base ##(0)%expr
- else
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if args0 =? 0
- then UnderLets.Base ##(0)%expr
- else
- if args =? 1
- then UnderLets.Base x0
- else
- if args0 =? 1
- then UnderLets.Base x
- else
- if args =? -1
- then UnderLets.Base (- x0)%expr
- else
- if args0 =? -1
- then UnderLets.Base (- x)%expr
- else
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
if args <? 0
- then UnderLets.Base (- (##((- args)%Z) * x0))%expr
- else
- if args0 <? 0
- then UnderLets.Base (- (x * ##((- args0)%Z)))%expr
- else
- if
- (args0 =? 2 ^ Z.log2 args0) && negb (args0 =? 2)
- then UnderLets.Base (x << ##(Z.log2 args0))%expr
- else
- if (args =? 2 ^ Z.log2 args) && negb (args =? 2)
- then UnderLets.Base (x0 << ##(Z.log2 args))%expr
- else UnderLets.Base (x * x0)%expr
- | None =>
- if args =? 1
- then UnderLets.Base x0
- else
- if args =? -1
- then UnderLets.Base (- x0)%expr
- else
- if args <? 0
- then UnderLets.Base (- (##((- args)%Z) * x0))%expr
- else
- if (args =? 2 ^ Z.log2 args) && negb (args =? 2)
- then UnderLets.Base (x0 << ##(Z.log2 args))%expr
- else UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s _ f x1 =>
- if args =? 1
- then UnderLets.Base x0
- else
- match f with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
+ then
+ Some
+ (UnderLets.Base (- (x' v + ##(- args)%Z))%expr)
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat
+ _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match x with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 >? 0
+ then Some (UnderLets.Base (##args0 - x' v)%expr)
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1);;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 <? 0
+ then
+ Some
+ (UnderLets.Base (- (##(- args0)%Z + x' v))%expr)
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x2 =>
+ _ <- pattern.ident.invert_bind_args idc0 pattern.ident.Z_opp;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
match
- s as t2
+ s as t3
return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
+ | type.base t3 =>
+ fun v0 : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
(fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? -1
- then UnderLets.Base (x' v)
- else
- if args =? -1
- then UnderLets.Base (- x0)%expr
- else
- if args <? 0
- then
- UnderLets.Base
- (- (##((- args)%Z) * x0))%expr
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- UnderLets.Base
- (- (x * x'0 v0))%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
+ a0 : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr
+ (type.base base.type.Z)) =>
+ match a0 with
+ | Some x'0 =>
+ Some
+ (UnderLets.Base
+ (- (x' v + x'0 v0))%expr)
+ | None => None
end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None =>
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 =>
+ None
+ end (reflect x1)
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2)
+ | @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' => Some (UnderLets.Base (x0 - x' v)%expr)
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' => Some (UnderLets.Base (x - x' v)%expr)
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end);;
+ None);;;
+ UnderLets.Base (x + x0)%expr)%option
+| ident.Z_mul =>
+ fun x x0 : defaults.expr (type.base base.type.Z) =>
+ (((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 0 then Some (UnderLets.Base (##0)%expr) else None)
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 0 then Some (UnderLets.Base (##0)%expr) else None)
+ | _ => None
+ end;;
+ match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 1 then Some (UnderLets.Base x0) else None)
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 1 then Some (UnderLets.Base x) else None);;
+ match x with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t0 idc0) x1 =>
+ _ <- pattern.ident.invert_bind_args idc0 pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
if args =? -1
- then UnderLets.Base (- x0)%expr
- else
- if args <? 0
- then UnderLets.Base (- (##((- args)%Z) * x0))%expr
- else
- if (args =? 2 ^ Z.log2 args) && negb (args =? 2)
- then UnderLets.Base (x0 << ##(Z.log2 args))%expr
- else UnderLets.Base (x * x0)%expr
- end
- | _ =>
- if args =? -1
- then UnderLets.Base (- x0)%expr
- else
- if args <? 0
- then UnderLets.Base (- (##((- args)%Z) * x0))%expr
- else
- if (args =? 2 ^ Z.log2 args) && negb (args =? 2)
- then UnderLets.Base (x0 << ##(Z.log2 args))%expr
- else UnderLets.Base (x * x0)%expr
- end
- | _ =>
- if args =? 1
- then UnderLets.Base x0
- else
- if args =? -1
- then UnderLets.Base (- x0)%expr
- else
- if args <? 0
- then UnderLets.Base (- (##((- args)%Z) * x0))%expr
- else
- if (args =? 2 ^ Z.log2 args) && negb (args =? 2)
- then UnderLets.Base (x0 << ##(Z.log2 args))%expr
- else UnderLets.Base (x * x0)%expr
- end
- | None =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base ##(0)%expr
- else
- if args =? 1
- then UnderLets.Base x
- else
- if args =? -1
- then UnderLets.Base (- x)%expr
- else
- if args <? 0
- then UnderLets.Base (- (x * ##((- args)%Z)))%expr
- else
- if (args =? 2 ^ Z.log2 args) && negb (args =? 2)
- then UnderLets.Base (x << ##(Z.log2 args))%expr
- else UnderLets.Base (x * x0)%expr
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s _ #(idc0)%expr_pat x1 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x * x' v))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => UnderLets.Base (x * x0)%expr
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | ($_)%expr =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t1 v0 =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base ##(0)%expr
- else
- if args =? 1
- then UnderLets.Base x
- else
- if args =? -1
- then UnderLets.Base (- x)%expr
- else
- if args <? 0
- then UnderLets.Base (- (x * ##((- args)%Z)))%expr
- else
- if (args =? 2 ^ Z.log2 args) && negb (args =? 2)
- then UnderLets.Base (x << ##(Z.log2 args))%expr
- else UnderLets.Base (x * x0)%expr
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s _ #(idc)%expr_pat x1 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1)) t2
- base.type.Z
+ then Some (UnderLets.Base (x' v))
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat
+ _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match x with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
(UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x * x' v0))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun _ : Compile.value' false s0 -> Compile.value' true d0
- => UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat
- _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x * x0)%expr
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base ##(0)%expr
- else
- if args =? 1
- then UnderLets.Base x
- else
- if args =? -1
- then UnderLets.Base (- x)%expr
- else
- if args <? 0
- then UnderLets.Base (- (x * ##((- args)%Z)))%expr
- else
- if (args =? 2 ^ Z.log2 args) && negb (args =? 2)
- then UnderLets.Base (x << ##(Z.log2 args))%expr
- else UnderLets.Base (x * x0)%expr
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s0 _ #(idc)%expr_pat x1 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 =? -1
+ then Some (UnderLets.Base (x' v))
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1)
+ | _ => None
+ end
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? -1 then Some (UnderLets.Base (- x0)%expr) else None)
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? -1 then Some (UnderLets.Base (- x)%expr) else None)
+ | _ => None
+ end;;
+ match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args <? 0
+ then Some (UnderLets.Base (- (##(- args)%Z * x0))%expr)
+ else None)
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args <? 0
+ then Some (UnderLets.Base (- (x * ##(- args)%Z))%expr)
+ else None)
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match x with
+ | @expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x2 =>
+ _ <- pattern.ident.invert_bind_args idc0 pattern.ident.Z_opp;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
(UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x * x' v))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1
- => UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0 _ (_ @ _)%expr_pat
- _ | @expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x * x0)%expr
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s _ f x1 =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base ##(0)%expr
- else
- if args =? 1
- then UnderLets.Base x
- else
- match f with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
match
- s as t2
+ s as t3
return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
+ | type.base t3 =>
+ fun v0 : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
(fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? -1
- then UnderLets.Base (x' v)
- else
- if args =? -1
- then UnderLets.Base (- x)%expr
- else
- if args <? 0
- then
- UnderLets.Base
- (- (x * ##((- args)%Z)))%expr
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- UnderLets.Base
- (- (x'0 v0 * x0))%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
+ a0 : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr
+ (type.base base.type.Z)) =>
+ match a0 with
+ | Some x'0 =>
+ Some
+ (UnderLets.Base (x' v * x'0 v0)%expr)
+ | None => None
end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None =>
- if args =? -1
- then UnderLets.Base (- x)%expr
- else
- if args <? 0
- then UnderLets.Base (- (x * ##((- args)%Z)))%expr
- else
- if (args =? 2 ^ Z.log2 args) && negb (args =? 2)
- then UnderLets.Base (x << ##(Z.log2 args))%expr
- else UnderLets.Base (x * x0)%expr
- end
- | _ =>
- if args =? -1
- then UnderLets.Base (- x)%expr
- else
- if args <? 0
- then UnderLets.Base (- (x * ##((- args)%Z)))%expr
- else
- if (args =? 2 ^ Z.log2 args) && negb (args =? 2)
- then UnderLets.Base (x << ##(Z.log2 args))%expr
- else UnderLets.Base (x * x0)%expr
- end
- | None =>
- match f with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- UnderLets.Base (- (x' v * x0))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | ($_)%expr =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x' v0 * x0))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x' v * x0))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s0 _ #(idc)%expr_pat x2 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match f with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- UnderLets.Base
- (x' v * x'0 v0)%expr
- | None =>
- UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- UnderLets.Base (- (x * x' v))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- end
- | ($_)%expr =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x * x' v0))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x * x' v))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x * x' v))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- end
- | None =>
- match f with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- UnderLets.Base (- (x' v * x0))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x' v0 * x0))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0
- _ (_ @ _)%expr_pat _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x' v * x0))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x' v * x0))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x' v * x0))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base ##(0)%expr
- else
- if args =? 1
- then UnderLets.Base x
- else
- if args =? -1
- then UnderLets.Base (- x)%expr
- else
- if args <? 0
- then UnderLets.Base (- (x * ##((- args)%Z)))%expr
- else
- if (args =? 2 ^ Z.log2 args) && negb (args =? 2)
- then UnderLets.Base (x << ##(Z.log2 args))%expr
- else UnderLets.Base (x * x0)%expr
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s _ #(idc)%expr_pat x2 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x * x' v))%expr
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun _ : Compile.value' false s0 -> Compile.value' true d0
- => UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat
- _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x * x0)%expr
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 =>
+ None
+ end (reflect x1)
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2)
+ | @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' => Some (UnderLets.Base (- (x' v * x0))%expr)
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if (args =? 2 ^ Z.log2 args) && negb (args =? 2)
+ then Some (UnderLets.Base (x << ##(Z.log2 args))%expr)
+ else None)
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' => Some (UnderLets.Base (- (x * x' v))%expr)
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if (args =? 2 ^ Z.log2 args) && negb (args =? 2)
+ then Some (UnderLets.Base (x0 << ##(Z.log2 args))%expr)
+ else None)
+ | _ => None
+ end);;
+ None);;;
+ UnderLets.Base (x * x0)%expr)%option
| ident.Z_pow =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat
| ident.Z_sub =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if args =? 0
- then UnderLets.Base (- x0)%expr
- else
- if args0 =? 0
- then UnderLets.Base x
- else
- if args <? 0
- then UnderLets.Base (- (##((- args)%Z) + x0))%expr
- else
- if args0 <? 0
- then UnderLets.Base (x + ##((- args0)%Z))%expr
- else UnderLets.Base (x - x0)%expr
- | None =>
- if args =? 0
- then UnderLets.Base (- x0)%expr
- else
- if args <? 0
- then UnderLets.Base (- (##((- args)%Z) + x0))%expr
- else UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s _ #(idc0)%expr_pat x1 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 0
- then UnderLets.Base (x' v)
- else
- if args =? 0
- then UnderLets.Base (- x0)%expr
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args >? 0
- then
- UnderLets.Base
- (##(args) + x'0 v0)%expr
- else
- match
- s as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v1 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr
- (type.base t1)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a1 with
- | Some x'1 =>
- if args <? 0
- then
- UnderLets.Base
- (x'1 v1 -
- ##((- args)%Z))%expr
- else
- if args <? 0
- then
- UnderLets.Base
- (-
- (##((- args)%Z) +
- x0))%expr
- else
- match
- s as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t1 : base.type
- =>
- defaults.expr
- (type.base
- t1)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- UnderLets.Base
- (x +
- x'2 v2)%expr
- | None =>
- UnderLets.Base
- (x - x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value'
- false s0 ->
- Compile.value'
- true d0
- =>
- UnderLets.Base
- (x - x0)%expr
- end
- (Compile.reflect
- x1)
- | None =>
- UnderLets.Base
- (x - x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false
- s0 ->
- Compile.value' true d0
- =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None =>
- UnderLets.Base (x - x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None =>
- if args =? 0
- then UnderLets.Base (- x0)%expr
- else
- if args <? 0
- then UnderLets.Base (- (##((- args)%Z) + x0))%expr
- else UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- if args =? 0
- then UnderLets.Base (- x0)%expr
- else
- if args <? 0
- then UnderLets.Base (- (##((- args)%Z) + x0))%expr
- else UnderLets.Base (x - x0)%expr
- | _ =>
- if args =? 0
- then UnderLets.Base (- x0)%expr
- else
- if args <? 0
- then UnderLets.Base (- (##((- args)%Z) + x0))%expr
- else UnderLets.Base (x - x0)%expr
- end
- | None =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base x
- else
- if args <? 0
- then UnderLets.Base (x + ##((- args)%Z))%expr
- else UnderLets.Base (x - x0)%expr
- | None => UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s _ #(idc0)%expr_pat x1 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x + x' v)%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => UnderLets.Base (x - x0)%expr
- | _ => UnderLets.Base (x - x0)%expr
- end
- end
- | ($_)%expr =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t1 v0 =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base x
- else
- if args <? 0
- then UnderLets.Base (x + ##((- args)%Z))%expr
- else UnderLets.Base (x - x0)%expr
- | None => UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s _ #(idc)%expr_pat x1 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1)) t2
- base.type.Z
+ (((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ match x0 with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t0 idc0) x1 =>
+ _ <- pattern.ident.invert_bind_args idc0 pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
(UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x + x' v0)%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun _ : Compile.value' false s0 -> Compile.value' true d0
- => UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat
- _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x - x0)%expr
- | _ => UnderLets.Base (x - x0)%expr
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base x
- else
- if args <? 0
- then UnderLets.Base (x + ##((- args)%Z))%expr
- else UnderLets.Base (x - x0)%expr
- | None => UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s0 _ #(idc)%expr_pat x1 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 0
+ then Some (UnderLets.Base (x' v))
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat
+ _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ (if args =? 0 then Some (UnderLets.Base (- x0)%expr) else None)
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 0 then Some (UnderLets.Base x) else None)
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match x with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
(UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x + x' v)%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1
- => UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0 _ (_ @ _)%expr_pat
- _ | @expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x - x0)%expr
- | _ => UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s _ f x1 =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base x
- else
- match f with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args >? 0
- then
- UnderLets.Base
- (- (x' v + ##((- args)%Z)))%expr
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.Base
- (##((- args)%Z) - x'0 v0)%expr
- else
- if args <? 0
- then
- UnderLets.Base
- (x + ##((- args)%Z))%expr
- else
- match
- s as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v1 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type
- =>
- defaults.expr
- (type.base t1)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a1 with
- | Some x'1 =>
- UnderLets.Base
- (-
- (x'1 v1 + x0))%expr
- | None =>
- UnderLets.Base
- (x - x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value'
- false s0 ->
- Compile.value'
- true d0 =>
- UnderLets.Base
- (x - x0)%expr
- end (Compile.reflect x1)
- | None =>
- UnderLets.Base (x - x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None =>
- if args <? 0
- then UnderLets.Base (x + ##((- args)%Z))%expr
- else UnderLets.Base (x - x0)%expr
- end
- | _ =>
- if args <? 0
- then UnderLets.Base (x + ##((- args)%Z))%expr
- else UnderLets.Base (x - x0)%expr
- end
- | None =>
- match f with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- UnderLets.Base (- (x' v + x0))%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x - x0)%expr
- end
- | _ => UnderLets.Base (x - x0)%expr
- end
- end
- | ($_)%expr =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x' v0 + x0))%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x - x0)%expr
- end
- | _ => UnderLets.Base (x - x0)%expr
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x' v + x0))%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x - x0)%expr
- end
- | _ => UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s0 _ #(idc)%expr_pat x2 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match f with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- UnderLets.Base
- (x'0 v0 - x' v)%expr
- | None =>
- UnderLets.Base (x - x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' => UnderLets.Base (x + x' v)%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x2)
- end
- | ($_)%expr =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x + x' v0)%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x2)
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x + x' v)%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x2)
- | _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x + x' v)%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x2)
- end
- | None =>
- match f with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- UnderLets.Base (- (x' v + x0))%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x - x0)%expr
- end
- | _ => UnderLets.Base (x - x0)%expr
- end
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x' v0 + x0))%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x - x0)%expr
- end
- | _ => UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0
- _ (_ @ _)%expr_pat _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x' v + x0))%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x - x0)%expr
- end
- | _ => UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x' v + x0))%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x - x0)%expr
- end
- | _ => UnderLets.Base (x - x0)%expr
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match f with
- | #(idc)%expr_pat =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (- (x' v + x0))%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base (x - x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x - x0)%expr
- end
- | _ => UnderLets.Base (x - x0)%expr
- end
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base x
- else
- if args <? 0
- then UnderLets.Base (x + ##((- args)%Z))%expr
- else UnderLets.Base (x - x0)%expr
- | None => UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s _ #(idc)%expr_pat x2 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 >? 0
+ then Some (UnderLets.Base (##args0 + x' v)%expr)
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1);;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
(UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x + x' v)%expr
- | None => UnderLets.Base (x - x0)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun _ : Compile.value' false s0 -> Compile.value' true d0
- => UnderLets.Base (x - x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x - x0)%expr
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat
- _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x - x0)%expr
- | _ => UnderLets.Base (x - x0)%expr
- end
- end
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 <? 0
+ then
+ Some (UnderLets.Base (x' v - ##(- args0)%Z)%expr)
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1)
+ | _ => None
+ end
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args <? 0
+ then Some (UnderLets.Base (- (##(- args)%Z + x0))%expr)
+ else None)
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 >? 0
+ then
+ Some
+ (UnderLets.Base (- (x' v + ##(- args0)%Z))%expr)
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1);;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 <? 0
+ then
+ Some (UnderLets.Base (##(- args0)%Z - x' v)%expr)
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1)
+ | _ => None
+ end
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args <? 0
+ then Some (UnderLets.Base (x + ##(- args)%Z)%expr)
+ else None)
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match x with
+ | @expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x2 =>
+ _ <- pattern.ident.invert_bind_args idc0 pattern.ident.Z_opp;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ match
+ s as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t3 =>
+ fun v0 : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
+ (fun
+ a0 : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr
+ (type.base base.type.Z)) =>
+ match a0 with
+ | Some x'0 =>
+ Some
+ (UnderLets.Base (x'0 v0 - x' v)%expr)
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 =>
+ None
+ end (reflect x1)
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2)
+ | @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' => Some (UnderLets.Base (- (x' v + x0))%expr)
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' => Some (UnderLets.Base (x + x' v)%expr)
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end);;
+ None);;;
+ UnderLets.Base (x - x0)%expr)%option
| ident.Z_opp =>
fun x : defaults.expr (type.base base.type.Z) =>
- match x with
- | ($_)%expr =>
- if negb (SubstVarLike.is_var_fst_snd_pair_opp x)
- then
- UnderLets.UnderLet x
- (fun v0 : var (type.base base.type.Z) =>
- UnderLets.Base (- $v0)%expr)
- else UnderLets.Base (- x)%expr
- | @expr.App _ _ _ s _ #(idc)%expr_pat x0 =>
- match match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
+ (((match x with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x0 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
(UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' => UnderLets.Base (x' v)
- | None => UnderLets.Base (- x)%expr
- end)
- | (s0 -> d0)%ptype =>
- fun _ : Compile.value' false s0 -> Compile.value' true d0 =>
- UnderLets.Base (- x)%expr
- end (Compile.reflect x0)
- | None =>
- if negb (SubstVarLike.is_var_fst_snd_pair_opp x)
- then
- UnderLets.UnderLet x
- (fun v : var (type.base base.type.Z) =>
- UnderLets.Base (- $v)%expr)
- else UnderLets.Base (- x)%expr
- end
- | @expr.App _ _ _ s _ ($_)%expr _ =>
- if negb (SubstVarLike.is_var_fst_snd_pair_opp x)
- then
- UnderLets.UnderLet x
- (fun v0 : var (type.base base.type.Z) =>
- UnderLets.Base (- $v0)%expr)
- else UnderLets.Base (- x)%expr
- | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _)
- _ =>
- if negb (SubstVarLike.is_var_fst_snd_pair_opp x)
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' => Some (UnderLets.Base (x' v))
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x0)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ (if negb (SubstVarLike.is_var_fst_snd_pair_opp x)
then
- UnderLets.UnderLet x
- (fun v : var (type.base base.type.Z) => UnderLets.Base (- $v)%expr)
- else UnderLets.Base (- x)%expr
- | _ =>
- if negb (SubstVarLike.is_var_fst_snd_pair_opp x)
- then
- UnderLets.UnderLet x
- (fun v : var (type.base base.type.Z) => UnderLets.Base (- $v)%expr)
- else UnderLets.Base (- x)%expr
- end
+ Some
+ (UnderLets.UnderLet x
+ (fun v : var (type.base base.type.Z) =>
+ UnderLets.Base (- $v)%expr))
+ else None));;
+ None);;;
+ UnderLets.Base (- x)%expr)%option
| ident.Z_div =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 1
- then UnderLets.Base x
- else
- if args =? 2 ^ Z.log2 args
- then UnderLets.Base (x >> ##(Z.log2 args))%expr
- else UnderLets.Base (x / x0)%expr
- | None => UnderLets.Base (x / x0)%expr
- end
- | _ => UnderLets.Base (x / x0)%expr
- end
+ ((match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 1 then Some (UnderLets.Base x) else None);;
+ (if args =? 2 ^ Z.log2 args
+ then Some (UnderLets.Base (x >> ##(Z.log2 args))%expr)
+ else None)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (x / x0)%expr)%option
| ident.Z_modulo =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 1
- then UnderLets.Base ##(0)%expr
- else
- if args =? 2 ^ Z.log2 args
- then UnderLets.Base (x &' ##((args - 1)%Z))%expr
- else UnderLets.Base (x mod x0)%expr
- | None => UnderLets.Base (x mod x0)%expr
- end
- | _ => UnderLets.Base (x mod x0)%expr
- end
+ (match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 1 then Some (UnderLets.Base (##0)%expr) else None);;
+ (if args =? 2 ^ Z.log2 args
+ then Some (UnderLets.Base (x &' ##(args - 1)%Z)%expr)
+ else None)
+ | _ => None
+ end;;;
+ UnderLets.Base (x mod x0)%expr)%option
| ident.Z_log2 =>
fun x : defaults.expr (type.base base.type.Z) =>
UnderLets.Base (#(ident.Z_log2)%expr @ x)%expr_pat
@@ -4163,125 +1265,20 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base (x << x0)%expr
| ident.Z_land =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base ##(0)%expr
- else
- match x with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if args0 =? 0
- then UnderLets.Base ##(0)%expr
- else UnderLets.Base (x &' x0)%expr
- | None => UnderLets.Base (x &' x0)%expr
- end
- | _ => UnderLets.Base (x &' x0)%expr
- end
- | None =>
- match x with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base ##(0)%expr
- else UnderLets.Base (x &' x0)%expr
- | None => UnderLets.Base (x &' x0)%expr
- end
- | _ => UnderLets.Base (x &' x0)%expr
- end
- end
- | ($_)%expr =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t1 v0 =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base ##(0)%expr
- else UnderLets.Base (x &' x0)%expr
- | None => UnderLets.Base (x &' x0)%expr
- end
- | _ => UnderLets.Base (x &' x0)%expr
- end
- | _ =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base ##(0)%expr
- else UnderLets.Base (x &' x0)%expr
- | None => UnderLets.Base (x &' x0)%expr
- end
- | _ => UnderLets.Base (x &' x0)%expr
- end
- end
+ (((match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 0 then Some (UnderLets.Base (##0)%expr) else None)
+ | _ => None
+ end;;
+ match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 0 then Some (UnderLets.Base (##0)%expr) else None)
+ | _ => None
+ end);;
+ None);;;
+ UnderLets.Base (x &' x0)%expr)%option
| ident.Z_lor =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
UnderLets.Base (x || x0)%expr
@@ -4293,3328 +1290,261 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base (#(ident.Z_lnot_modulo)%expr @ x @ x0)%expr_pat
| ident.Z_mul_split =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some _ =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if args0 =? 0
- then UnderLets.Base (##(0)%expr, ##(0)%expr)%expr_pat
- else
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- if args1 =? 0
- then
- UnderLets.Base
- (##(0)%expr, ##(0)%expr)%expr_pat
- else
- if args0 =? 1
- then UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- if args1 =? 1
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- if args0 =? -1
- then
- UnderLets.Base
- ((- x1)%expr, ##(0)%expr)%expr_pat
- else
- if args1 =? -1
- then
- UnderLets.Base
- ((- x0)%expr, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args0 =? 1
- then UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- if args0 =? -1
- then
- UnderLets.Base
- ((- x1)%expr, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args0 =? 1
- then UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- if args0 =? -1
- then
- UnderLets.Base ((- x1)%expr, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args0 =? 1
- then UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- if args0 =? -1
- then
- UnderLets.Base ((- x1)%expr, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if args0 =? 0
- then
- UnderLets.Base (##(0)%expr, ##(0)%expr)%expr_pat
- else
- if args0 =? 1
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- if args0 =? -1
- then
- UnderLets.Base
- ((- x0)%expr, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | ($_)%expr =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t2 v0 =>
- match
- t2 as t3 return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args0 =>
- if args0 =? 0
- then UnderLets.Base (##(0)%expr, ##(0)%expr)%expr_pat
- else
- if args0 =? 1
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- if args0 =? -1
- then
- UnderLets.Base ((- x0)%expr, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | _ =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if args0 =? 0
- then UnderLets.Base (##(0)%expr, ##(0)%expr)%expr_pat
- else
- if args0 =? 1
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- if args0 =? -1
- then
- UnderLets.Base ((- x0)%expr, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr, #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr, #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
+ (((match x with
+ | @expr.Ident _ _ _ t idc =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ (if args0 =? 0
+ then Some (UnderLets.Base ((##0)%expr, (##0)%expr)%expr_pat)
+ else None)
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ (if args0 =? 0
+ then Some (UnderLets.Base ((##0)%expr, (##0)%expr)%expr_pat)
+ else None)
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ (if args0 =? 1
+ then Some (UnderLets.Base (x1, (##0)%expr)%expr_pat)
+ else None)
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ (if args0 =? 1
+ then Some (UnderLets.Base (x0, (##0)%expr)%expr_pat)
+ else None)
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ (if args0 =? -1
+ then Some (UnderLets.Base ((- x1)%expr, (##0)%expr)%expr_pat)
+ else None)
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ (if args0 =? -1
+ then Some (UnderLets.Base ((- x0)%expr, (##0)%expr)%expr_pat)
+ else None)
+ | _ => None
+ end
+ | _ => None
+ end;;
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
+ (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
+ UnderLets.Base
+ (#(ident.fst)%expr @ ($v)%expr, #(ident.snd)%expr @ ($v)%expr)%expr_pat)));;
+ None);;;
+ UnderLets.Base (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat)%option
| ident.Z_add_get_carry =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x1 @
- ##((- args)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if args0 <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- ##((- args0)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- if args0 =? 0
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x1 @
- ##((- args)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x1 @
- ##((- args)%Z)%expr)%expr_pat
- (fun
- v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s _ #(idc0)%expr_pat x2 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x2)
- | None =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x1 @
- ##((- args)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ ($_)%expr _ =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x1 @
- ##((- args)%Z)%expr)%expr_pat
- (fun
- v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _
- s _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x1 @
- ##((- args)%Z)%expr)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x1 @
- ##((- args)%Z)%expr)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- ##((- args)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s _ #(idc0)%expr_pat x2 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x2)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _
- s _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | ($_)%expr =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t1 v0 =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- ##((- args)%Z)%expr)%expr_pat
- (fun
- v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v1 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | @expr.App _ _ _ s _ #(idc)%expr_pat x2 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x' v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- (- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun _ : Compile.value' false s0 -> Compile.value' true d0
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x2)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v1 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- ##((- args)%Z)%expr)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ #(idc)%expr_pat x2 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x2)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0
- _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ #(idc)%expr_pat x2 =>
- match match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
+ (((match x0 with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x2 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
(UnderLets.UnderLets base.type ident var
(defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x1 @ x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun _ : Compile.value' false s0 -> Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x2)
- | None =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- ##((- args)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ #(idc0)%expr_pat x3 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
- _ s0 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | @expr.App _ _ _ s _ ($_)%expr _ =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t1 v0 =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- ##((- args)%Z)%expr)%expr_pat
- (fun
- v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v1 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ #(idc)%expr_pat x3 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x' v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- (- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v1 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0
- _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- ##((- args)%Z)%expr)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ #(idc)%expr_pat x3 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s1 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun _ : Compile.value' false s2 -> Compile.value' true d2
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s1 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s1
- _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s1 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- ##((- args)%Z)%expr)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ #(idc)%expr_pat x4 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s1 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun _ : Compile.value' false s2 -> Compile.value' true d2
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x4)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s1 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s1
- _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s1 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- ##((- args)%Z)%expr)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ #(idc)%expr_pat x4 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x4)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0
- _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- ##((- args)%Z)%expr)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then UnderLets.Base (x0, ##(0)%expr)%expr_pat
- else
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun
- v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s _ #(idc)%expr_pat x3 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun _ : Compile.value' false s0 -> Compile.value' true d0
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
-| ident.Z_add_with_carry =>
- fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- if args =? 0
- then UnderLets.Base (x0 + x1)%expr
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base base.type.Z) =>
- UnderLets.Base ($v)%expr)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base base.type.Z) =>
- UnderLets.Base ($v)%expr)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v0 : var (type.base base.type.Z) => UnderLets.Base ($v0)%expr)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (type.base base.type.Z) => UnderLets.Base ($v)%expr)
- end
-| ident.Z_add_with_get_carry =>
- fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- if
- (args0 <=? 0) && (args <=? 0) &&
- (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x2 @
- ##((- args0)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- if
- (args1 <=? 0) && (args <=? 0) &&
- (args1 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x1 @
- ##((- args1)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base
- (x2, ##(0)%expr)%expr_pat
- else
- if (args =? 0) && (args1 =? 0)
- then
- UnderLets.Base
- (x1, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- if (args0 =? 0) && (args1 =? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ ##(args0)%expr @
- ##(args1)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- ##(0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- if (args0 =? 0) && (args1 =? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ ##(args0)%expr @
- ##(args1)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- ##(0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- if (args0 =? 0) && (args1 =? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ ##(args0)%expr @ ##(args1)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- ##(0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- if (args0 =? 0) && (args1 =? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ ##(args0)%expr @ ##(args1)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- ##(0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- if
- (args0 <=? 0) && (args <=? 0) &&
- (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x2 @
- ##((- args0)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base
- (x2, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | ($_)%expr =>
- if
- (args0 <=? 0) && (args <=? 0) && (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x2 @ ##((- args0)%Z)%expr)%expr_pat
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x1 @ x' v)%expr_pat
(fun
v0 : var
(type.base
(base.type.Z * base.type.Z)%etype) =>
UnderLets.Base
(#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base (x2, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ #(idc1)%expr_pat x3 =>
- match
- match idc1 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @
- x1 @ x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @
- ##((- args)%Z)%expr @
- x1 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- else
- if
- (args0 <=? 0) &&
- (args <=? 0) &&
- (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @
- ##((- args)%Z)%expr @
- x2 @
- ##((- args0)%Z)%expr)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal
- t5 v1 =>
- match
- t5 as t6
- return
- (base.base_interp
- t6 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun v2 : Z
- =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ => None
- end
- with
- | Some _ =>
- if
- (args =? 0) &&
- (args0 =? 0)
- then
- UnderLets.Base
- (x2,
- ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 :
- var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (
- #
- (ident.fst)%expr @
- ($v1)%expr,
- #
- (ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v1 :
- var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (
- #
- (ident.fst)%expr @
- ($v1)%expr,
- #
- (ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 :
- var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#
- (ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v1 :
- var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#
- (ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- #(ident.snd)%expr @
- ($v2)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- #(ident.snd)%expr @
- ($v2)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if
- (args0 <=? 0) && (args <=? 0) &&
- (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x2 @
- ##((- args0)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base
- (x2, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | @expr.App _ _ _ s _ ($_)%expr _ =>
- if
- (args0 <=? 0) && (args <=? 0) && (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x2 @ ##((- args0)%Z)%expr)%expr_pat
+ (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat))
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x2)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x2 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x' v)%expr_pat
(fun
v0 : var
(type.base
(base.type.Z * base.type.Z)%etype) =>
UnderLets.Base
(#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base (x2, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s _ (_ @ _)%expr_pat _ | @expr.App _ _
- _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- if
- (args0 <=? 0) && (args <=? 0) && (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x2 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base (x2, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
+ (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat))
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x2)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args <? 0
+ then
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x1 @
+ (##(- args)%Z)%expr)%expr_pat
+ (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
+ =>
+ UnderLets.Base
+ (#(ident.fst)%expr @ ($v)%expr,
+ (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat))
+ else None)
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args <? 0
+ then
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
+ (##(- args)%Z)%expr)%expr_pat
+ (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
+ =>
+ UnderLets.Base
+ (#(ident.fst)%expr @ ($v)%expr,
+ (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat))
+ else None)
+ | _ => None
+ end;;
+ match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 0
+ then Some (UnderLets.Base (x1, (##0)%expr)%expr_pat)
+ else None)
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 0
+ then Some (UnderLets.Base (x0, (##0)%expr)%expr_pat)
+ else None)
+ | _ => None
+ end;;
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
+ (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
+ UnderLets.Base
+ (#(ident.fst)%expr @ ($v)%expr, #(ident.snd)%expr @ ($v)%expr)%expr_pat)));;
+ None);;;
+ UnderLets.Base (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat)%option
+| ident.Z_add_with_carry =>
+ fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
+ (((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ (if args =? 0 then Some (UnderLets.Base (x0 + x1)%expr) else None)
+ | _ => None
+ end;;
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
+ (fun v : var (type.base base.type.Z) => UnderLets.Base ($v)%expr)));;
+ None);;;
+ UnderLets.Base (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat)%option
+| ident.Z_add_with_get_carry =>
+ fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
+ (((match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ match x1 with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t0 idc0) x3 =>
+ _ <- pattern.ident.invert_bind_args idc0 pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 0
+ then
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x2 @
+ x' v)%expr_pat
(fun
v0 : var
(type.base
@@ -7622,11 +1552,37 @@ match idc in (ident t) return (Compile.value' true t) with
=>
UnderLets.Base
(#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
+ (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat))
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x3);;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args <? 0
+ then
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @
+ (##(- args)%Z)%expr @ x2 @ x' v)%expr_pat
(fun
v0 : var
(type.base
@@ -7634,127 +1590,45 @@ match idc in (ident t) return (Compile.value' true t) with
=>
UnderLets.Base
(#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | _ =>
- if
- (args0 <=? 0) && (args <=? 0) && (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x2 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base (x2, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
+ (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat))
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x3)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat
+ _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x2 with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t0 idc0) x3 =>
+ _ <- pattern.ident.invert_bind_args idc0 pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 0
+ then
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x1 @
+ x' v)%expr_pat
(fun
v0 : var
(type.base
@@ -7762,11 +1636,37 @@ match idc in (ident t) return (Compile.value' true t) with
=>
UnderLets.Base
(#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
+ (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat))
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x3);;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args <? 0
+ then
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @
+ (##(- args)%Z)%expr @ x1 @ x' v)%expr_pat
(fun
v0 : var
(type.base
@@ -7774,4594 +1674,267 @@ match idc in (ident t) return (Compile.value' true t) with
=>
UnderLets.Base
(#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
+ (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat))
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x3)
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat
+ _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ (if (args0 <=? 0) && (args <=? 0) && (args0 + args <? 0)
+ then
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @
+ (##(- args)%Z)%expr @ x2 @ (##(- args0)%Z)%expr)%expr_pat
+ (fun
+ v : var
+ (type.base (base.type.Z * base.type.Z)%etype)
+ =>
+ UnderLets.Base
+ (#(ident.fst)%expr @ ($v)%expr,
+ (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat))
+ else None)
+ | _ => None
+ end;;
+ match x2 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ (if (args0 <=? 0) && (args <=? 0) && (args0 + args <? 0)
+ then
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @
+ (##(- args)%Z)%expr @ x1 @ (##(- args0)%Z)%expr)%expr_pat
+ (fun
+ v : var
+ (type.base (base.type.Z * base.type.Z)%etype)
+ =>
+ UnderLets.Base
+ (#(ident.fst)%expr @ ($v)%expr,
+ (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat))
+ else None)
+ | _ => None
+ end;;
+ match x with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ _ <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ (if (args =? 0) && (args1 =? 0)
+ then Some (UnderLets.Base (x2, (##0)%expr)%expr_pat)
+ else None)
+ | _ => None
+ end;;
+ match x2 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ (if (args =? 0) && (args1 =? 0)
+ then Some (UnderLets.Base (x1, (##0)%expr)%expr_pat)
+ else None)
+ | _ => None
+ end
+ | _ => None
+ end;;
+ (if args =? 0
+ then
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
+ (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
+ =>
+ UnderLets.Base
+ (#(ident.fst)%expr @ ($v)%expr,
+ #(ident.snd)%expr @ ($v)%expr)%expr_pat))
+ else None)
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x3 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.Z_opp;
+ match x1 with
+ | @expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x4 =>
+ _ <- pattern.ident.invert_bind_args idc0 pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ match
+ s0 as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base
+ (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t3 =>
+ fun v0 : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
(fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | None =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if
- (args0 <=? 0) && (args <=? 0) &&
- (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x1 @
- ##((- args0)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base
- (x1, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s _ #(idc1)%expr_pat x3 =>
- match
- match idc1 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @
- x1 @ x' v)%expr_pat
+ a0 : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr
+ (type.base base.type.Z)) =>
+ match a0 with
+ | Some x'0 =>
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_with_get_borrow)%expr @
+ x @ x' v @ x2 @ x'0 v0)%expr_pat
(fun
- v0 : var
+ v1 : var
(type.base
(base.type.Z *
base.type.Z)%etype)
=>
UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
+ (#(ident.fst)%expr @ ($v1)%expr,
(-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @
- ##((- args)%Z)%expr @
- x1 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ ($_)%expr _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s _ (_ @ _)%expr_pat _ | @expr.App _ _
- _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | ($_)%expr =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t2 v0 =>
- match
- t2 as t3 return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args0 =>
- if
- (args0 <=? 0) && (args <=? 0) && (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | @expr.App _ _ _ s _ #(idc0)%expr_pat x3 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @
- x1 @ x' v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- (-
- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v1 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ ##((- args)%Z)%expr @
- x1 @ x'0 v1)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- (-
- (#(ident.snd)%expr @
- $v2)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- #(ident.snd)%expr @
- ($v2)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- #(ident.snd)%expr @
- ($v2)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ ($_)%expr _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _
- _ _ s _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
+ (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat))
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 =>
+ None
+ end (reflect x4)
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x3)
+ | @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x2 with
+ | @expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x4 =>
+ _ <- pattern.ident.invert_bind_args idc0 pattern.ident.Z_opp;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
match
- t1 as t2 return (base.base_interp t2 -> option Z)
+ s0 as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base
+ (base.type.Z * base.type.Z)%etype))))
with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if
- (args0 <=? 0) && (args <=? 0) && (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
+ | type.base t3 =>
+ fun v0 : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
(fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ #(idc0)%expr_pat x3 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @
- x1 @ x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ ##((- args)%Z)%expr @
- x1 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s0 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ #(idc0)%expr_pat x3 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @ x2 @
- x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ ##((- args)%Z)%expr @ x2 @
- x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- else
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t4 v1 =>
- match
- t4 as t5
- return
- (base.base_interp
- t5 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit =>
- None
- | base.type.Z =>
- fun v2 : Z =>
- Some v2
- | base.type.bool =>
- fun _ : bool =>
- None
- | base.type.nat =>
- fun _ : nat => None
- end v1
- | _ => None
- end
- with
- | Some args1 =>
- if
- (args1 <=? 0) &&
- (args <=? 0) &&
- (args1 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @
- ##((- args)%Z)%expr @
- x1 @
- ##((- args1)%Z)%expr)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal
- t5 v1 =>
- match
- t5 as t6
- return
- (base.base_interp
- t6 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun v2 : Z
- =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ => None
- end
- with
- | Some _ =>
- if
- (args =? 0) &&
- (args1 =? 0)
- then
- UnderLets.Base
- (x1,
- ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 :
- var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (
- #
- (ident.fst)%expr @
- ($v1)%expr,
- #
- (ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v1 :
- var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (
- #
- (ident.fst)%expr @
- ($v1)%expr,
- #
- (ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 :
- var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#
- (ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v1 :
- var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#
- (ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- #(ident.snd)%expr @
- ($v2)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- #(ident.snd)%expr @
- ($v2)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- end
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- #(ident.snd)%expr @
- ($v2)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- #(ident.snd)%expr @
- ($v2)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _
- #(idc1)%expr_pat x4 =>
- match
- match idc1 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v1 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a1 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a1 with
- | Some x'1 =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @
- x @ x1 @
- x'1 v1)%expr_pat
- (fun
- v2 :
- var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (
- #
- (ident.fst)%expr @
- ($v2)%expr,
- (-
- (#
- (ident.snd)%expr @
- $v2)%expr_pat)%expr)%expr_pat)
- else
- match
- s0 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base
- t5 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args <? 0
- then
- UnderLets.UnderLet
- (#
- (ident.Z_sub_with_get_borrow)%expr @
- x @
- ##
- ((- args)%Z)%expr @
- x1 @
- x'2 v2)%expr_pat
- (fun
- v3 :
- var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (
- #
- (ident.fst)%expr @
- ($v3)%expr,
- (-
- (#
- (ident.snd)%expr @
- $v3)%expr_pat)%expr)%expr_pat)
- else
- if
- args =? 0
- then
- UnderLets.UnderLet
- (#
- (ident.Z_add_get_carry)%expr @
- x @ x1 @
- x2)%expr_pat
- (fun
- v3 :
- var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (
- #
- (ident.fst)%expr @
- ($v3)%expr,
- #
- (ident.snd)%expr @
- ($v3)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v3 :
- var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (
- #
- (ident.fst)%expr @
- ($v3)%expr,
- #
- (ident.snd)%expr @
- ($v3)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value'
- true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ ($_)%expr
- _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- #(ident.snd)%expr @
- ($v2)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- #(ident.snd)%expr @
- ($v2)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _
- (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s0 _
- (_ @ _)%expr_pat _ | @expr.App
- _ _ _ s0 _
- (@expr.LetIn _ _ _ _ _ _ _)
- _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if
- (args0 <=? 0) && (args <=? 0) &&
- (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x1 @
- ##((- args0)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base
- (x1, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ #(idc1)%expr_pat x4 =>
- match
- match idc1 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @
- x1 @ x' v)%expr_pat
+ a0 : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr
+ (type.base base.type.Z)) =>
+ match a0 with
+ | Some x'0 =>
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_with_get_borrow)%expr @
+ x @ x' v @ x1 @ x'0 v0)%expr_pat
(fun
- v0 : var
+ v1 : var
(type.base
(base.type.Z *
base.type.Z)%etype)
=>
UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
+ (#(ident.fst)%expr @ ($v1)%expr,
(-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @
- ##((- args)%Z)%expr @
- x1 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s0 _ (_ @ _)%expr_pat _ | @expr.App _ _
- _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | @expr.App _ _ _ s _ ($_)%expr _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t2 v0 =>
- match
- t2 as t3 return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args0 =>
- if
- (args0 <=? 0) && (args <=? 0) && (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ #(idc0)%expr_pat x4 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @
- x1 @ x' v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- (-
- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- else
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v1 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ ##((- args)%Z)%expr @
- x1 @ x'0 v1)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- (-
- (#(ident.snd)%expr @
- $v2)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- #(ident.snd)%expr @
- ($v2)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- #(ident.snd)%expr @
- ($v2)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s0 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if
- (args0 <=? 0) && (args <=? 0) && (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ #(idc0)%expr_pat x4 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s1 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @
- x1 @ x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ ##((- args)%Z)%expr @
- x1 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s1 _ ($_)%expr _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s1 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s1 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if
- (args0 <=? 0) && (args <=? 0) && (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ #(idc0)%expr_pat x5 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s1 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @
- x1 @ x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ ##((- args)%Z)%expr @
- x1 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s1 _ ($_)%expr _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s1 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s1 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if
- (args0 <=? 0) && (args <=? 0) && (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ #(idc0)%expr_pat x5 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @
- x1 @ x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ ##((- args)%Z)%expr @
- x1 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s0 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if
- (args0 <=? 0) && (args <=? 0) && (args0 + args <? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- ##((- args)%Z)%expr @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- (- (#(ident.snd)%expr @ $v)%expr_pat)%expr)%expr_pat)
- else
- match x with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some _ =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.Base (x1, ##(0)%expr)%expr_pat
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
+ (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat))
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 =>
+ None
+ end (reflect x4)
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x3)
+ | @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 <=? 0
+ then
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @
+ x' v @ x2 @ (##(- args0)%Z)%expr)%expr_pat
(fun
v0 : var
(type.base
@@ -12369,11 +1942,43 @@ match idc in (ident t) return (Compile.value' true t) with
=>
UnderLets.Base
(#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
+ (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat))
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x3)
+ | _ => None
+ end;;
+ match x2 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 <=? 0
+ then
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @
+ x' v @ x1 @ (##(- args0)%Z)%expr)%expr_pat
(fun
v0 : var
(type.base
@@ -12381,4287 +1986,53 @@ match idc in (ident t) return (Compile.value' true t) with
=>
UnderLets.Base
(#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s _ #(idc0)%expr_pat x4 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_get_borrow)%expr @ x @
- x1 @ x' v)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args <? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ ##((- args)%Z)%expr @
- x1 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- else
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ ($_)%expr _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _
- _ _ s _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | None =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- ##(args)%expr @ ##(args0)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr, ##(0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | ($_)%expr =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t1 v0 =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t2 v0 =>
- match
- t2 as t3 return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args0 =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- ##(args)%expr @ ##(args0)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr, ##(0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v1 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ #(idc)%expr_pat x3 =>
- match match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end with
- | Some _ =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x2 @
- ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args1 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x'0 v0 @ x1 @
- ##((- args1)%Z)%expr)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- else
- if
- (args0 =? 0) &&
- (args1 =? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- ##(args0)%expr @
- ##(args1)%expr)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- ##(0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- #(ident.snd)%expr @
- ($v1)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x2 @
- ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- end
- | ($_)%expr =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v0 @ x2 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- (-
- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | @expr.Abs _ _ _ _ _ _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x2 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | @expr.App _ _ _ s0 _ #(idc1)%expr_pat x4 =>
- match
- match idc1 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a0 with
- | Some x'0 =>
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @
- x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x2 @
- ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v0 @ x2 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- (-
- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s0 _ (_ @ _)%expr_pat _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x2 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | @expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x2 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x2 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- end
- | None =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @
- ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ #(idc1)%expr_pat x4 =>
- match
- match idc1 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a0 with
- | Some x'0 =>
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @
- x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s0 _ (_ @ _)%expr_pat _ | @expr.App _ _
- _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | ($_)%expr =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t2 v0 =>
- match
- t2 as t3 return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args0 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v0 @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- (-
- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ #(idc0)%expr_pat x4 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v1 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v0 @ x1 @ x'0 v1)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- (-
- (#(ident.snd)%expr @ $v2)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s0 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ #(idc0)%expr_pat x4 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s1 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s1 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s1 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ #(idc0)%expr_pat x4 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a0 with
- | Some x'0 =>
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x2 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @
- ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ #(idc1)%expr_pat x5 =>
- match
- match idc1 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a0 with
- | Some x'0 =>
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @
- x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s1 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s1 _ (_ @ _)%expr_pat _ | @expr.App _ _
- _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t2 v0 =>
- match
- t2 as t3 return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args0 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v0 @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- (-
- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ #(idc0)%expr_pat x5 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v1 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v0 @ x1 @ x'0 v1)%expr_pat
- (fun
- v2 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v2)%expr,
- (-
- (#(ident.snd)%expr @ $v2)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s1 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s1 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s1 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s2 _ #(idc0)%expr_pat x5 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s2 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s2 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s2 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ (_ @ _)%expr_pat _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s2 _ #(idc0)%expr_pat x6 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s2 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s2 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s2 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ #(idc0)%expr_pat x6 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s1 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s1 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s1 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 <=? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @ ##((- args0)%Z)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- (-
- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ #(idc0)%expr_pat x5 =>
- match
- match idc0 with
- | ident.Z_opp => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- UnderLets.UnderLet
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x' v @ x1 @ x'0 v0)%expr_pat
- (fun
- v1 : var
- (type.base
- (base.type.Z *
- base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @
- ($v1)%expr,
- (-
- (#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | @expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s0 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | None =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- ##(args)%expr @ ##(args0)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr, ##(0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
- | @expr.App _ _ _ s _ ($_)%expr _ =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t1 v0 =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t2 v0 =>
- match
- t2 as t3 return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args0 =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- ##(args)%expr @ ##(args0)%expr)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr, ##(0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v0 : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v1 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v1 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v1)%expr,
- #(ident.snd)%expr @ ($v1)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- end
- | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _)
- _ =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- ##(args)%expr @ ##(args0)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr, ##(0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | _ =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match
- t0 as t1 return (base.base_interp t1 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if (args =? 0) && (args0 =? 0)
- then
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- ##(args)%expr @ ##(args0)%expr)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr, ##(0)%expr)%expr_pat)
- else
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- (fun
- v : var
- (type.base
- (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun
- v0 : var
- (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
+ (- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat))
+ else None
+ | None => None
+ end)
+ | (s0 -> d0)%ptype =>
+ fun _ : value' false s0 -> value' true d0 => None
+ end (reflect x3)
+ | _ => None
+ end
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ match x2 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ (if (args =? 0) && (args0 =? 0)
+ then
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
+ (##args)%expr @ (##args0)%expr)%expr_pat
(fun
v : var
(type.base (base.type.Z * base.type.Z)%etype)
=>
UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | None =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype)
- =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- | ($_)%expr =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v0 : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v0)%expr,
- #(ident.snd)%expr @ ($v0)%expr)%expr_pat)
- | _ =>
- UnderLets.UnderLet
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
- UnderLets.Base
- (#(ident.fst)%expr @ ($v)%expr,
- #(ident.snd)%expr @ ($v)%expr)%expr_pat)
- end
- end
+ (#(ident.fst)%expr @ ($v)%expr, (##0)%expr)%expr_pat))
+ else None)
+ | _ => None
+ end
+ | _ => None
+ end;;
+ Some
+ (UnderLets.UnderLet
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
+ (fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
+ UnderLets.Base
+ (#(ident.fst)%expr @ ($v)%expr, #(ident.snd)%expr @ ($v)%expr)%expr_pat)));;
+ None);;;
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat)%option
| ident.Z_sub_get_borrow =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
UnderLets.UnderLet
@@ -16693,100 +2064,74 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base (#(ident.Z_cast range)%expr @ x)%expr_pat
| ident.Z_cast2 range =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- match x with
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x1) x0 =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
+ (match x with
+ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc) x1) x0 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.pair;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
(defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ match
+ s as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
(defaults.expr
- (type.base (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t0 : base.type =>
- defaults.expr (type.base t0)) t3 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a0 : option
- (defaults.expr (type.base t3) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- (fv <-- do_again
- (base.type.Z * base.type.Z)
- (#(ident.Z_cast (fst range))%expr @
- ($(x' v))%expr,
- #(ident.Z_cast (snd range))%expr @
- ($(x'0 v0))%expr)%expr_pat;
- UnderLets.Base (id (id fv)))%under_lets
- | None =>
- UnderLets.Base
- (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- end (Compile.reflect x0)
- | None =>
- UnderLets.Base
- (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1 =>
- UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- end (Compile.reflect x1)
- | None => UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- end
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App _
- _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ | @expr.App
- _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ | @expr.App _ _ _
- s _ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ =>
- UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr _ |
- @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- | _ => UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- end
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t3 =>
+ fun v0 : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
+ (fun
+ a0 : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr (type.base base.type.Z))
+ =>
+ match a0 with
+ | Some x'0 =>
+ Some
+ (fv <-- do_again (base.type.Z * base.type.Z)
+ (#(ident.Z_cast (fst range))%expr @
+ ($(x' v))%expr,
+ #(ident.Z_cast (snd range))%expr @
+ ($(x'0 v0))%expr)%expr_pat;
+ UnderLets.Base (id (id fv)))%under_lets
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x0)
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App _
+ _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ | @expr.App
+ _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ | @expr.App _ _
+ _ s _ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ => None
+ | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr
+ _ | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s
+ _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;;
+ UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat)%option
| ident.fancy_add log2wordmax imm =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
UnderLets.Base (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
@@ -16839,4 +2184,4 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat
end
- : Compile.value' true t
+ : value' true t
diff --git a/src/Experiments/NewPipeline/fancy_rewrite_head.out b/src/Experiments/NewPipeline/fancy_rewrite_head.out
index b93038b7c..0dfac41db 100644
--- a/src/Experiments/NewPipeline/fancy_rewrite_head.out
+++ b/src/Experiments/NewPipeline/fancy_rewrite_head.out
@@ -1,5 +1,5 @@
fancy_rewrite_head =
-match idc in (ident t) return (Compile.value' true t) with
+match idc in (ident t) return (value' true t) with
| @ident.Literal t v =>
match
t as t0
@@ -8,10 +8,10 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base t0)))
with
- | base.type.unit => fun v0 : unit => UnderLets.Base ##(v0)%expr
- | base.type.Z => fun v0 : Z => UnderLets.Base ##(v0)%expr
- | base.type.bool => fun v0 : bool => UnderLets.Base ##(v0)%expr
- | base.type.nat => fun v0 : nat => UnderLets.Base ##(v0)%expr
+ | base.type.unit => fun v0 : unit => UnderLets.Base (##v0)%expr
+ | base.type.Z => fun v0 : Z => UnderLets.Base (##v0)%expr
+ | base.type.bool => fun v0 : bool => UnderLets.Base (##v0)%expr
+ | base.type.nat => fun v0 : nat => UnderLets.Base (##v0)%expr
end v
| ident.Nat_succ =>
fun x : defaults.expr (type.base base.type.nat) =>
@@ -235,17457 +235,1247 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base (x + x0)%expr
| ident.Z_mul =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ (match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ match x0 with
+ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x2) x1 =>
+ (_ <- pattern.ident.invert_bind_args idc0 pident.Z_land;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args1 =? 2 ^ (2 * Z.log2_up args1 / 2) - 1
+ then
+ (x3 <- invert_low (2 * Z.log2_up args1) args;
+ Some
+ (Some
+ (UnderLets.Base
+ (#(ident.fancy_mulll
+ (2 * Z.log2_up args1))%expr @
+ ((##x3)%expr, x' v))%expr_pat)));;;
+ None
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2)
+ | _ => None
+ end;;
+ match x2 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args1 =? 2 ^ (2 * Z.log2_up args1 / 2) - 1
+ then
+ (x3 <- invert_low (2 * Z.log2_up args1) args;
+ Some
+ (Some
+ (UnderLets.Base
+ (#(ident.fancy_mulll
+ (2 * Z.log2_up args1))%expr @
+ ((##x3)%expr, x' v))%expr_pat)));;;
+ None
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x1);;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args1 =? 2 ^ (2 * Z.log2_up args1 / 2) - 1
+ then
+ (x3 <- invert_high (2 * Z.log2_up args1) args;
+ Some
+ (Some
+ (UnderLets.Base
+ (#(ident.fancy_mulhl
+ (2 * Z.log2_up args1))%expr @
+ ((##x3)%expr, x' v))%expr_pat)));;;
+ None
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x1)
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args1 =? 2 ^ (2 * Z.log2_up args1 / 2) - 1
+ then
+ (x3 <- invert_high (2 * Z.log2_up args1) args;
+ Some
+ (Some
+ (UnderLets.Base
+ (#(ident.fancy_mulhl
+ (2 * Z.log2_up args1))%expr @
+ ((##x3)%expr, x' v))%expr_pat)));;;
+ None
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2)
+ | _ => None
+ end);;
+ _ <- pattern.ident.invert_bind_args idc0 pident.Z_shiftr;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ (x3 <- invert_low (2 * args1) args;
+ Some
+ (Some
+ (UnderLets.Base
+ (#(ident.fancy_mullh (2 * args1))%expr @
+ ((##x3)%expr, x' v))%expr_pat)));;;
+ None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2);;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ (x3 <- invert_high (2 * args1) args;
+ Some
+ (Some
+ (UnderLets.Base
+ (#(ident.fancy_mulhh (2 * args1))%expr @
+ ((##x3)%expr, x' v))%expr_pat)));;;
+ None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2)
+ | _ => None
+ end
+ | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ |
+ @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ | @expr.App _ _
+ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ | @expr.App _ _
+ _ s _ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ =>
+ None
+ | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _
+ ($_)%expr _ | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end
+ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc) x2) x1 =>
+ (_ <- pattern.ident.invert_bind_args idc pident.Z_land;
+ match x2 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 =? 2 ^ (2 * Z.log2_up args0 / 2) - 1
+ then
+ (y <- invert_low (2 * Z.log2_up args0) args1;
+ Some
+ (Some
+ (UnderLets.Base
+ (#(ident.fancy_mulll
+ (2 * Z.log2_up args0))%expr @
+ (x' v, (##y)%expr))%expr_pat)));;;
+ None
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x1)
+ | _ => None
+ end
| _ => None
- end
- with
- | Some args =>
- match x0 with
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc0)%expr_pat x2)
- x1 =>
- match
- match idc0 with
- | ident.Z_land => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if
- args1 =?
- 2 ^ (2 * Z.log2_up args1 / 2) - 1
- then
- match
- invert_low (2 * Z.log2_up args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 * Z.log2_up args1))%expr @
- (##(x3)%expr, x' v))%expr_pat
- | None =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp
- t5 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit =>
- None
- | base.type.Z =>
- fun v1 : Z =>
- Some v1
- | base.type.bool =>
- fun _ : bool =>
- None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args2 =?
- 2
- ^ (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_low
- (2 *
- Z.log2_up
- args2)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args2))%expr @
- (##
- (x3)%expr,
- x'0 v0))%expr_pat
- | None =>
- match
- s as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t5 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- if
- args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args2)
- args
- with
- | Some
- x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args2))%expr @
- (##
- (x3)%expr,
- x'1 v1))%expr_pat
- | None =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args1 =?
- 2
- ^
- (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some
- x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'2 v2))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- end
- else
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args1 =?
- 2
- ^
- (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some
- x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'2 v2))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x1)
- end
- else
- match
- s as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t5 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- if
- args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args2)
- args
- with
- | Some
- x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args2))%expr @
- (##
- (x3)%expr,
- x'1 v1))%expr_pat
- | None =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args1 =?
- 2
- ^
- (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some
- x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'2 v2))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- end
- else
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args1 =?
- 2
- ^
- (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some
- x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'2 v2))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x1)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value'
- true d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x1)
- | None =>
- match
- s0 as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value'
- true d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- end
- | ($_)%expr =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v1 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'0 v1))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | @expr.Abs _ _ _ _ _ _ =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value'
- false s2 ->
- Compile.value' true
- d2 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | (_ @ _)%expr_pat =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some x4 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x4)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value'
- false s2 ->
- Compile.value' true
- d2 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some x4 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x4)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- end
- end
- else
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args2 =?
- 2
- ^ (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_low
- (2 *
- Z.log2_up
- args2)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 *
- Z.log2_up
- args2))%expr @
- (##
- (x3)%expr,
- x'0 v0))%expr_pat
- | None =>
- match
- s as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t5 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- if
- args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args2)
- args
- with
- | Some
- x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args2))%expr @
- (##
- (x3)%expr,
- x'1 v1))%expr_pat
- | None =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args1 =?
- 2
- ^
- (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some
- x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'2 v2))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- end
- else
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args1 =?
- 2
- ^
- (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some
- x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'2 v2))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x1)
- end
- else
- match
- s as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- if
- args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args2)
- args
- with
- | Some
- x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args2))%expr @
- (##
- (x3)%expr,
- x'1 v1))%expr_pat
- | None =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args1 =?
- 2
- ^
- (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some
- x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'2 v2))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- end
- else
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args1 =?
- 2
- ^
- (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some
- x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'2 v2))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x1)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x1)
- | None =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- end
- | ($_)%expr =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v1 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up args1 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##(x3)%expr,
- x'0 v1))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | @expr.Abs _ _ _ _ _ _ =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up args1 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##(x3)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | (_ @ _)%expr_pat =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up args1 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args1)
- args
- with
- | Some x4 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##(x4)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up args1 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args1)
- args
- with
- | Some x4 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##(x4)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- end
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 =? 2 ^ (2 * Z.log2_up args0 / 2) - 1
+ then
+ (y <- invert_low (2 * Z.log2_up args0) args1;
+ Some
+ (Some
+ (UnderLets.Base
+ (#(ident.fancy_mulll
+ (2 * Z.log2_up args0))%expr @
+ (x' v, (##y)%expr))%expr_pat)));;;
+ None
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2)
+ | _ => None
+ end
+ | _ => None
+ end;;
+ match x2 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 =? 2 ^ (2 * Z.log2_up args0 / 2) - 1
+ then
+ (y <- invert_high (2 * Z.log2_up args0) args1;
+ Some
+ (Some
+ (UnderLets.Base
+ (#(ident.fancy_mullh
+ (2 * Z.log2_up args0))%expr @
+ (x' v, (##y)%expr))%expr_pat)));;;
+ None
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x1)
+ | _ => None
+ end
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 =? 2 ^ (2 * Z.log2_up args0 / 2) - 1
+ then
+ (y <- invert_high (2 * Z.log2_up args0) args1;
+ Some
+ (Some
+ (UnderLets.Base
+ (#(ident.fancy_mullh
+ (2 * Z.log2_up args0))%expr @
+ (x' v, (##y)%expr))%expr_pat)));;;
+ None
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2)
+ | _ => None
+ end
+ | _ => None
+ end;;
+ match x2 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match x0 with
+ | @expr.App _ _ _ s1 _
+ (@expr.Ident _ _ _ t1 idc1 @ x4)%expr_pat x3 =>
+ _ <- pattern.ident.invert_bind_args idc1 pident.Z_land;
+ match x4 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pattern.ident.LiteralZ;
+ match
+ s as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t3 =>
+ fun v : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr (type.base base.type.Z))
+ =>
+ match a with
+ | Some x' =>
+ match
+ s1 as t4
+ return
+ (value' false t4 ->
+ option
+ (UnderLets.UnderLets base.type ident
var
(defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3)
+ (type.base base.type.Z))))
+ with
+ | type.base t4 =>
+ fun v0 : defaults.expr (type.base t4) =>
+ base.try_make_transport_cps t4
+ base.type.Z
+ (fun
+ a0 : option
+ (defaults.expr
+ (type.base t4) ->
+ defaults.expr
+ (type.base base.type.Z))
=>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args1 =?
- 2 ^ (2 * Z.log2_up args1 / 2) -
- 1
- then
- match
- invert_low
- (2 * Z.log2_up args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 *
- Z.log2_up args1))%expr @
- (##(x3)%expr, x' v))%expr_pat
- | None =>
- match
- s as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value'
- true d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x1)
- end
- else
- match
- s as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x1)
- | None =>
- UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | ($_)%expr =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if
- args1 =?
- 2 ^ (2 * Z.log2_up args1 / 2) - 1
- then
- match
- invert_low
- (2 * Z.log2_up args1) args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 * Z.log2_up args1))%expr @
- (##(x3)%expr, x' v0))%expr_pat
- | None =>
- match
- s as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v1 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'0 v1))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x1)
- end
- else
- match
- s as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v1 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up args1 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##(x3)%expr,
- x'0 v1))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if
- args1 =?
- 2 ^ (2 * Z.log2_up args1 / 2) - 1
- then
- match
- invert_low
- (2 * Z.log2_up args1) args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 * Z.log2_up args1))%expr @
- (##(x3)%expr, x' v))%expr_pat
- | None =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x3)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value'
- false s2 ->
- Compile.value' true
- d2 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x1)
- end
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up args1 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##(x3)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | (_ @ _)%expr_pat =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if
- args1 =?
- 2 ^ (2 * Z.log2_up args1 / 2) - 1
- then
- match
- invert_low
- (2 * Z.log2_up args1) args
- with
- | Some x4 =>
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 * Z.log2_up args1))%expr @
- (##(x4)%expr, x' v))%expr_pat
- | None =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some x4 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x4)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value'
- false s2 ->
- Compile.value' true
- d2 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x1)
- end
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up args1 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args1)
- args
- with
- | Some x4 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##(x4)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if
- args1 =?
- 2 ^ (2 * Z.log2_up args1 / 2) - 1
- then
- match
- invert_low
- (2 * Z.log2_up args1) args
- with
- | Some x4 =>
- UnderLets.Base
+ match a0 with
+ | Some x'0 =>
+ if
+ (args0 =?
+ 2 ^ (2 * Z.log2_up args0 / 2) -
+ 1) &&
+ (args2 =?
+ 2 ^ (2 * Z.log2_up args0 / 2) -
+ 1)
+ then
+ Some
+ (UnderLets.Base
(#(ident.fancy_mulll
- (2 * Z.log2_up args1))%expr @
- (##(x4)%expr, x' v))%expr_pat
- | None =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up
- args1 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args1)
- args
- with
- | Some x4 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##
- (x4)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x1)
- end
- else
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args1 =?
- 2
- ^ (2 *
- Z.log2_up args1 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args1)
- args
- with
- | Some x4 =>
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- Z.log2_up
- args1))%expr @
- (##(x4)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- match
- invert_low (2 * args1) args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 * args1))%expr @
- (##(x3)%expr, x' v))%expr_pat
- | None =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- match
- invert_high
- (2 * args1)
- args
- with
- | Some x3 =>
- UnderLets.Base
- (#(ident.fancy_mulhh
- (2 *
- args1))%expr @
- (##(x3)%expr,
- x'0 v0))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false
- s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- end
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- end
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ |
- @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ | @expr.App
- _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ |
- @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ =>
- UnderLets.Base (x * x0)%expr
- | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _
- ($_)%expr _ | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x * x0)%expr
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x2) x1 =>
- match match idc with
- | ident.Z_land => Some tt
+ (2 * Z.log2_up args0))%expr @
+ (x' v, x'0 v0))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun
+ _ : value' false s3 -> value' true d3
+ => None
+ end (reflect x3)
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun _ : value' false s3 -> value' true d3 => None
+ end (reflect x1)
+ | _ => None
+ end
+ | @expr.App _ _ _ s1 _ #(_)%expr_pat _ | @expr.App _ _ _ s1 _
+ ($_)%expr _ | @expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _)
+ _ | @expr.App _ _ _ s1 _ (($_)%expr @ _)%expr_pat _ |
+ @expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat _ |
+ @expr.App _ _ _ s1 _ (_ @ _ @ _)%expr_pat _ | @expr.App _ _ _
+ s1 _ (@expr.LetIn _ _ _ _ _ _ _ @ _)%expr_pat _ | @expr.App _
+ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
- end with
- | Some _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
+ end
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match x0 with
+ | @expr.App _ _ _ s1 _
+ (@expr.Ident _ _ _ t1 idc1 @ x4)%expr_pat x3 =>
+ _ <- pattern.ident.invert_bind_args idc1 pident.Z_land;
+ match x4 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pattern.ident.LiteralZ;
match
- t1 as t2 return (base.base_interp t2 -> option Z)
+ s0 as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if
- args0 =?
- 2 ^ (2 * Z.log2_up args0 / 2) - 1
- then
- match
- invert_low (2 * Z.log2_up args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 * Z.log2_up args0))%expr @
- (x' v, ##(y)%expr))%expr_pat
- | None =>
- match x1 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp
- t5 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit =>
- None
- | base.type.Z =>
- fun v1 : Z =>
- Some v1
- | base.type.bool =>
- fun _ : bool =>
- None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args2 =?
- 2
- ^ (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_low
- (2 *
- Z.log2_up
- args2)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args2))%expr @
- (
- x'0 v0,
- ##
- (y)%expr))%expr_pat
- | None =>
- match
- s as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t5 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- if
- args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some
- y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'1 v1,
- ##
- (y)%expr))%expr_pat
- | None =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args2)
- args1
- with
- | Some
- y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args2))%expr @
- (
- x'2 v2,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- end
- else
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args2)
- args1
- with
- | Some
- y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args2))%expr @
- (
- x'2 v2,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x1)
- end
- else
- match
- s as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t5 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- if
- args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some
- y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'1 v1,
- ##
- (y)%expr))%expr_pat
- | None =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args2)
- args1
- with
- | Some
- y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args2))%expr @
- (
- x'2 v2,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- end
- else
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args2)
- args1
- with
- | Some
- y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args2))%expr @
- (
- x'2 v2,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x1)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value'
- true d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- match
- s as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'0 v0,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value'
- true d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x1)
- end
- | ($_)%expr =>
- match
- s as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v1 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v1,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x1)
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v0,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x1)
- | _ =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v0,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value'
- false s2 ->
- Compile.value' true
- d2 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x1)
- end
- end
- else
- match x1 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args2 =?
- 2
- ^ (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_low
- (2 *
- Z.log2_up
- args2)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 *
- Z.log2_up
- args2))%expr @
- (x'0 v0,
- ##
- (y)%expr))%expr_pat
- | None =>
- match
- s as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t5 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- if
- args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some
- y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'1 v1,
- ##
- (y)%expr))%expr_pat
- | None =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args2)
- args1
- with
- | Some
- y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args2))%expr @
- (
- x'2 v2,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- end
- else
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args2)
- args1
- with
- | Some
- y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args2))%expr @
- (
- x'2 v2,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x1)
- end
- else
- match
- s as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- if
- args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some
- y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'1 v1,
- ##
- (y)%expr))%expr_pat
- | None =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args2)
- args1
- with
- | Some
- y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args2))%expr @
- (
- x'2 v2,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- end
- else
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args2 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args2)
- args1
- with
- | Some
- y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args2))%expr @
- (
- x'2 v2,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x1)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- match
- s as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v0,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x1)
- end
- | ($_)%expr =>
- match
- s as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v1 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up args0 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v1,
- ##(y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up args0 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v0,
- ##(y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | _ =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up args0 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v0,
- ##(y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- end
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ #(idc1)%expr_pat x4) x3 =>
- match
- match idc1 with
- | ident.Z_land => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
+ | type.base t3 =>
+ fun v : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr (type.base base.type.Z))
+ =>
+ match a with
+ | Some x' =>
+ match
+ s1 as t4
+ return
+ (value' false t4 ->
+ option
+ (UnderLets.UnderLets base.type ident
var
(defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s1 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (x' v, x'0 v0))%expr_pat
- else
- match x1 with
- | #(idc3)%expr_pat =>
- match
- match
- idc3
- with
- | @ident.Literal
- t6 v1 =>
- match
- t6 as t7
- return
- (base.base_interp
- t7 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v2 : Z =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ => None
- end
- with
- | Some
- args3 =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s1 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args3 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args3 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args3))%expr @
- (
- x'1 v1,
- x'2 v2))%expr_pat
- else
- match
- x3
- with
- | #
- (idc4)%expr_pat =>
- match
- match
- idc4
- with
- | @ident.Literal
- t9 v3 =>
- match
- t9 as t10
- return
- (base.base_interp
- t10 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v4 : Z =>
- Some v4
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v3
- | _ =>
- None
- end
- with
- | Some
- args4 =>
- match
- s as t9
- return
- (Compile.value'
- false t9 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t9 =>
- fun
- v3 :
- defaults.expr
- (type.base
- t9) =>
- base.try_make_transport_cps
- (fun
- t10 : base.type
- =>
- defaults.expr
- (type.base
- t10)) t9
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a3 :
- option
- (defaults.expr
- (type.base
- t9) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a3
- with
- | Some
- x'3 =>
- match
- s2 as t10
- return
- (Compile.value'
- false t10 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t10 =>
- fun
- v4 :
- defaults.expr
- (type.base
- t10) =>
- base.try_make_transport_cps
- (fun
- t11 : base.type
- =>
- defaults.expr
- (type.base
- t11)) t10
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a4 :
- option
- (defaults.expr
- (type.base
- t10) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a4
- with
- | Some
- x'4 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args4 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'3 v3,
- x'4 v4))%expr_pat
- else
- match
- s0 as t11
- return
- (Compile.value'
- false t11 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t11 =>
- fun
- v5 :
- defaults.expr
- (type.base
- t11) =>
- base.try_make_transport_cps
- (fun
- t12 : base.type
- =>
- defaults.expr
- (type.base
- t12)) t11
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a5 :
- option
- (defaults.expr
- (type.base
- t11) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a5
- with
- | Some
- x'5 =>
- match
- s2 as t12
- return
- (Compile.value'
- false t12 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t12 =>
- fun
- v6 :
- defaults.expr
- (type.base
- t12) =>
- base.try_make_transport_cps
- (fun
- t13 : base.type
- =>
- defaults.expr
- (type.base
- t13)) t12
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a6 :
- option
- (defaults.expr
- (type.base
- t12) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a6
- with
- | Some
- x'6 =>
- if
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args3 / 2) -
- 1) &&
- (args4 =?
- 2
- ^
- (2 *
- Z.log2_up
- args3 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args3))%expr @
- (
- x'5 v5,
- x'6 v6))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x1)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x3)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- match
- x3
- with
- | #
- (idc4)%expr_pat =>
- match
- match
- idc4
- with
- | @ident.Literal
- t7 v1 =>
- match
- t7 as t8
- return
- (base.base_interp
- t8 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v2 : Z =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s2 as t8
- return
- (Compile.value'
- false t8 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t8 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t8) =>
- base.try_make_transport_cps
- (fun
- t9 : base.type
- =>
- defaults.expr
- (type.base
- t9)) t8
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t8) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'1 v1,
- x'2 v2))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x1)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- end
- | ($_)%expr =>
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match
- idc3
- with
- | @ident.Literal
- t7 v2 =>
- match
- t7 as t8
- return
- (base.base_interp
- t8 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v3 : Z =>
- Some v3
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v2
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s2 as t8
- return
- (Compile.value'
- false t8 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t8 =>
- fun
- v3 :
- defaults.expr
- (type.base
- t8) =>
- base.try_make_transport_cps
- (fun
- t9 : base.type
- =>
- defaults.expr
- (type.base
- t9)) t8
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t8) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'1 v2,
- x'2 v3))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x1)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | @expr.LetIn _ _
- _ _ _ _ _ =>
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match
- idc3
- with
- | @ident.Literal
- t6 v1 =>
- match
- t6 as t7
- return
- (base.base_interp
- t7 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v2 : Z =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s2 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'1 v1,
- x'2 v2))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x1)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match
- idc3
- with
- | @ident.Literal
- t6 v1 =>
- match
- t6 as t7
- return
- (base.base_interp
- t7 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v2 : Z =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s2 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'1 v1,
- x'2 v2))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 ->
- d4)%ptype =>
- fun
- _ :
- Compile.value'
- false s4 ->
- Compile.value'
- true d4
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 ->
- d4)%ptype =>
- fun
- _ :
- Compile.value'
- false s4 ->
- Compile.value'
- true d4
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x1)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false
- s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None =>
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v0 : defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base t6))
- t5 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- match
- x1
- with
- | #
- (idc4)%expr_pat =>
- match
- match
- idc4
- with
- | @ident.Literal
- t7 v1 =>
- match
- t7 as t8
- return
- (base.base_interp
- t8 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v2 : Z =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s0 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s2 as t8
- return
- (Compile.value'
- false t8 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t8 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t8) =>
- base.try_make_transport_cps
- (fun
- t9 : base.type
- =>
- defaults.expr
- (type.base
- t9)) t8
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t8) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args3 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args3 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args3))%expr @
- (
- x'1 v1,
- x'2 v2))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | ($_)%expr =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s2 as t5
- return
- (Compile.value' false
- t5 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v1 : defaults.expr
- (type.base t5)
- =>
- base.try_make_transport_cps
- (fun t6 : base.type
- =>
- defaults.expr
- (type.base t6))
- t5 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (x' v0,
- x'0 v1))%expr_pat
- else
- match x1 with
- | #(idc3)%expr_pat =>
- match
- match
- idc3
- with
- | @ident.Literal
- t7 v2 =>
- match
- t7 as t8
- return
- (base.base_interp
- t8 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v3 : Z =>
- Some v3
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v2
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s0 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s2 as t8
- return
- (Compile.value'
- false t8 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t8 =>
- fun
- v3 :
- defaults.expr
- (type.base
- t8) =>
- base.try_make_transport_cps
- (fun
- t9 : base.type
- =>
- defaults.expr
- (type.base
- t9)) t8
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t8) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args3 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args3 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args3))%expr @
- (
- x'1 v2,
- x'2 v3))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s2 as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (x' v,
- x'0 v0))%expr_pat
- else
- match x1 with
- | #(idc3)%expr_pat =>
- match
- match
- idc3
- with
- | @ident.Literal
- t6 v1 =>
- match
- t6 as t7
- return
- (base.base_interp
- t7 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v2 : Z =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s2 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args3 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args3 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args3))%expr @
- (
- x'1 v1,
- x'2 v2))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
+ (type.base base.type.Z))))
+ with
+ | type.base t4 =>
+ fun v0 : defaults.expr (type.base t4) =>
+ base.try_make_transport_cps t4
+ base.type.Z
+ (fun
+ a0 : option
(defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s2 as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (x' v,
- x'0 v0))%expr_pat
- else
- match x1 with
- | #(idc3)%expr_pat =>
- match
- match
- idc3
- with
- | @ident.Literal
- t6 v1 =>
- match
- t6 as t7
- return
- (base.base_interp
- t7 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v2 : Z =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s2 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args3 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args3 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args3))%expr @
- (
- x'1 v1,
- x'2 v2))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 ->
- d4)%ptype =>
- fun
- _ :
- Compile.value'
- false s4 ->
- Compile.value'
- true d4
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 ->
- d4)%ptype =>
- fun
- _ :
- Compile.value'
- false s4 ->
- Compile.value'
- true d4
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value'
- true d4 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s2 as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 * args2 /
- 2) - 1
- then
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- args2))%expr @
- (x' v,
- x'0 v0))%expr_pat
- else
- match x1 with
- | #(idc3)%expr_pat =>
- match
- match
- idc3
- with
- | @ident.Literal
- t6 v1 =>
- match
- t6 as t7
- return
- (base.base_interp
- t7 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v2 : Z =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s2 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args3 =?
- 2
- ^
- (2 *
- args2 / 2) -
- 1
- then
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- args2))%expr @
- (
- x'1 v1,
- x'2 v2))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x1)
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- end
- | @expr.App _ _ _ s1 _ (@expr.App _ _ _ s2 _ ($_)%expr _)
- _ | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _) _ |
- @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ (_ @ _)%expr_pat _) _ | @expr.App
- _ _ _ s1 _
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _) _)
- _ => UnderLets.Base (x * x0)%expr
- | @expr.App _ _ _ s1 _ #(_)%expr_pat _ | @expr.App _ _ _
- s1 _ ($_)%expr _ | @expr.App _ _ _ s1 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s1 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x * x0)%expr
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3)
+ (type.base t4) ->
+ defaults.expr
+ (type.base base.type.Z))
=>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args0 =?
- 2 ^ (2 * Z.log2_up args0 / 2) -
- 1
- then
- match
- invert_low
- (2 * Z.log2_up args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 *
- Z.log2_up args0))%expr @
- (x' v, ##(y)%expr))%expr_pat
- | None =>
- match
- s0 as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'0 v0,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value'
- true d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- end
- else
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v0,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ #(idc2)%expr_pat x4)
- x3 =>
- match
- match idc2 with
- | ident.Z_land => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s1 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v0 : defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base t6))
- t5 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- match
- x3
- with
- | #
- (idc4)%expr_pat =>
- match
- match
- idc4
- with
- | @ident.Literal
- t7 v1 =>
- match
- t7 as t8
- return
- (base.base_interp
- t8 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v2 : Z =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s0 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s2 as t8
- return
- (Compile.value'
- false t8 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t8 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t8) =>
- base.try_make_transport_cps
- (fun
- t9 : base.type
- =>
- defaults.expr
- (type.base
- t9)) t8
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t8) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'1 v1,
- x'2 v2))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- match x3 with
- | #(idc4)%expr_pat =>
- match
- match idc4 with
- | @ident.Literal t5 v =>
- match
- t5 as t6
- return
- (base.base_interp
- t6 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit =>
- None
- | base.type.Z =>
- fun v0 : Z =>
- Some v0
- | base.type.bool =>
- fun _ : bool =>
- None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t5
- return
- (Compile.value' false
- t5 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v : defaults.expr
- (type.base t5)
- =>
- base.try_make_transport_cps
- (fun t6 : base.type
- =>
- defaults.expr
- (type.base t6))
- t5 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base (x * x0)%expr
- end
- end
- | ($_)%expr =>
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t5 v0 =>
- match
- t5 as t6
- return
- (base.base_interp t6 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t5
- return
- (Compile.value' false t5 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v0 : defaults.expr
- (type.base t5)
- =>
- base.try_make_transport_cps
- (fun t6 : base.type =>
- defaults.expr
- (type.base t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t6 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v0,
- x'0 v1))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | _ =>
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ :
- Compile.value'
- false s4 ->
- Compile.value'
- true d4
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value' true
- d4 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | None =>
- match
- match idc2 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- args0 =?
- 2
- ^
- (2 *
- args2 / 2) -
- 1
- then
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- args2))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- end
- | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ ($_)%expr _) _ |
- @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _)
- _) _ | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ (_ @ _)%expr_pat _) _ |
- @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _
- (@expr.LetIn _ _ _ _ _ _ _) _) _ =>
- UnderLets.Base (x * x0)%expr
- | @expr.App _ _ _ s1 _ #(_)%expr_pat _ |
- @expr.App _ _ _ s1 _ ($_)%expr _ | @expr.App _
- _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x * x0)%expr
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | ($_)%expr =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t2 v0 =>
- match
- t2 as t3 return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if
- args0 =?
- 2 ^ (2 * Z.log2_up args0 / 2) - 1
- then
- match
- invert_low
- (2 * Z.log2_up args0) args1
- with
- | Some y =>
- UnderLets.Base
+ match a0 with
+ | Some x'0 =>
+ if
+ (args0 =?
+ 2 ^ (2 * Z.log2_up args0 / 2) -
+ 1) &&
+ (args2 =?
+ 2 ^ (2 * Z.log2_up args0 / 2) -
+ 1)
+ then
+ Some
+ (UnderLets.Base
(#(ident.fancy_mulll
(2 * Z.log2_up args0))%expr @
- (x' v0, ##(y)%expr))%expr_pat
- | None =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v1 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v1,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- end
- else
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v1 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up args0 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v1,
- ##(y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ #(idc1)%expr_pat x4) x3 =>
- match
- match idc1 with
- | ident.Z_land => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
+ (x' v, x'0 v0))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun
+ _ : value' false s3 -> value' true d3
+ => None
+ end (reflect x3)
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun _ : value' false s3 -> value' true d3 => None
+ end (reflect x2)
+ | _ => None
+ end
+ | @expr.App _ _ _ s1 _ #(_)%expr_pat _ | @expr.App _ _ _ s1 _
+ ($_)%expr _ | @expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _)
+ _ | @expr.App _ _ _ s1 _ (($_)%expr @ _)%expr_pat _ |
+ @expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat _ |
+ @expr.App _ _ _ s1 _ (_ @ _ @ _)%expr_pat _ | @expr.App _ _ _
+ s1 _ (@expr.LetIn _ _ _ _ _ _ _ @ _)%expr_pat _ | @expr.App _
+ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end
+ | _ => None
+ end;;
+ match x2 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match x0 with
+ | (@expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t1 idc1) x4 @ x3)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc1 pident.Z_land;
+ match x3 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pattern.ident.LiteralZ;
+ match
+ s as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t3 =>
+ fun v : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr (type.base base.type.Z))
+ =>
+ match a with
+ | Some x' =>
+ match
+ s2 as t4
+ return
+ (value' false t4 ->
+ option
+ (UnderLets.UnderLets base.type ident
+ var
+ (defaults.expr
+ (type.base base.type.Z))))
+ with
+ | type.base t4 =>
+ fun v0 : defaults.expr (type.base t4) =>
+ base.try_make_transport_cps t4
+ base.type.Z
+ (fun
+ a0 : option
(defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s1 as t5
- return
- (Compile.value' false
- t5 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v1 : defaults.expr
- (type.base t5)
- =>
- base.try_make_transport_cps
- (fun t6 : base.type
- =>
- defaults.expr
- (type.base t6))
- t5 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (x' v0,
- x'0 v1))%expr_pat
- else
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match
- idc3
- with
- | @ident.Literal
- t7 v2 =>
- match
- t7 as t8
- return
- (base.base_interp
- t8 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v3 : Z =>
- Some v3
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v2
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s0 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s2 as t8
- return
- (Compile.value'
- false t8 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t8 =>
- fun
- v3 :
- defaults.expr
- (type.base
- t8) =>
- base.try_make_transport_cps
- (fun
- t9 : base.type
- =>
- defaults.expr
- (type.base
- t9)) t8
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t8) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'1 v2,
- x'2 v3))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t5 v0 =>
- match
- t5 as t6
- return
- (base.base_interp t6 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t5
- return
- (Compile.value' false t5 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v0 : defaults.expr
- (type.base t5)
- =>
- base.try_make_transport_cps
- (fun t6 : base.type =>
- defaults.expr
- (type.base t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t6 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v0,
- x'0 v1))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | ($_)%expr =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t5 v1 =>
- match
- t5 as t6
- return
- (base.base_interp t6 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v2 : Z => Some v2
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v1
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t5
- return
- (Compile.value' false t5 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t5 =>
- fun
- v1 : defaults.expr
- (type.base t5) =>
- base.try_make_transport_cps
- (fun t6 : base.type =>
- defaults.expr
- (type.base t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t6 =>
- fun
- v2 : defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base t7))
- t6 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v1,
- x'0 v2))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v1 : defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base t6))
- t5 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v0,
- x'0 v1))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v1 : defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base t6))
- t5 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v0,
- x'0 v1))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value'
- true d4 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v1 : defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base t6))
- t5 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^
- (2 * args2 /
- 2) - 1
- then
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- args2))%expr @
- (
- x' v0,
- x'0 v1))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- end
- | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ ($_)%expr _) _ | @expr.App _
- _ _ s1 _
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _)
- _ | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ (_ @ _)%expr_pat _) _ |
- @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _)
- _) _ => UnderLets.Base (x * x0)%expr
- | @expr.App _ _ _ s1 _ #(_)%expr_pat _ | @expr.App _
- _ _ s1 _ ($_)%expr _ | @expr.App _ _ _ s1 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s1 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x * x0)%expr
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if
- args0 =?
- 2 ^ (2 * Z.log2_up args0 / 2) - 1
- then
- match
- invert_low
- (2 * Z.log2_up args0) args1
- with
- | Some y =>
- UnderLets.Base
+ (type.base t4) ->
+ defaults.expr
+ (type.base base.type.Z))
+ =>
+ match a0 with
+ | Some x'0 =>
+ if
+ (args0 =?
+ 2 ^ (2 * Z.log2_up args0 / 2) -
+ 1) &&
+ (args2 =?
+ 2 ^ (2 * Z.log2_up args0 / 2) -
+ 1)
+ then
+ Some
+ (UnderLets.Base
(#(ident.fancy_mulll
(2 * Z.log2_up args0))%expr @
- (x' v, ##(y)%expr))%expr_pat
- | None =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v0,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value'
- false s2 ->
- Compile.value' true
- d2 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- end
- else
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up args0 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v0,
- ##(y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s2 _
- (@expr.App _ _ _ s3 _ #(idc1)%expr_pat x4) x3 =>
- match
- match idc1 with
- | ident.Z_land => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
+ (x' v, x'0 v0))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun
+ _ : value' false s3 -> value' true d3
+ => None
+ end (reflect x4)
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun _ : value' false s3 -> value' true d3 => None
+ end (reflect x1)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
+ | _ => None
+ end
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match x0 with
+ | (@expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t1 idc1) x4 @ x3)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc1 pident.Z_land;
+ match x3 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pattern.ident.LiteralZ;
+ match
+ s0 as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t3 =>
+ fun v : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr (type.base base.type.Z))
+ =>
+ match a with
+ | Some x' =>
+ match
+ s2 as t4
+ return
+ (value' false t4 ->
+ option
+ (UnderLets.UnderLets base.type ident
+ var
+ (defaults.expr
+ (type.base base.type.Z))))
+ with
+ | type.base t4 =>
+ fun v0 : defaults.expr (type.base t4) =>
+ base.try_make_transport_cps t4
+ base.type.Z
+ (fun
+ a0 : option
(defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s2 as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (x' v,
- x'0 v0))%expr_pat
- else
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match
- idc3
- with
- | @ident.Literal
- t6 v1 =>
- match
- t6 as t7
- return
- (base.base_interp
- t7 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v2 : Z =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s3 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'1 v1,
- x'2 v2))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 ->
- d4)%ptype =>
- fun
- _ :
- Compile.value'
- false s4 ->
- Compile.value'
- true d4
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 ->
- d4)%ptype =>
- fun
- _ :
- Compile.value'
- false s4 ->
- Compile.value'
- true d4
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value'
- true d4 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s3 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ :
- Compile.value'
- false s4 ->
- Compile.value'
- true d4
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value' true
- d4 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | ($_)%expr =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s3 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v1 : defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base t6))
- t5 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v0,
- x'0 v1))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value'
- true d4 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s3 as t4
- return
- (Compile.value'
- false t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun
- t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value'
- true d4 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s3 as t4
- return
- (Compile.value'
- false t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun
- t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s5 -> d5)%ptype =>
- fun
- _ : Compile.value'
- false s5 ->
- Compile.value'
- true d5 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s5 -> d5)%ptype =>
- fun
- _ : Compile.value' false s5 ->
- Compile.value' true d5
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s3 as t4
- return
- (Compile.value'
- false t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun
- t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^
- (2 * args2 /
- 2) - 1
- then
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- args2))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value'
- true d4 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- end
- | @expr.App _ _ _ s2 _
- (@expr.App _ _ _ s3 _ ($_)%expr _) _ | @expr.App _
- _ _ s2 _
- (@expr.App _ _ _ s3 _ (@expr.Abs _ _ _ _ _ _) _)
- _ | @expr.App _ _ _ s2 _
- (@expr.App _ _ _ s3 _ (_ @ _)%expr_pat _) _ |
- @expr.App _ _ _ s2 _
- (@expr.App _ _ _ s3 _ (@expr.LetIn _ _ _ _ _ _ _)
- _) _ => UnderLets.Base (x * x0)%expr
- | @expr.App _ _ _ s2 _ #(_)%expr_pat _ | @expr.App _
- _ _ s2 _ ($_)%expr _ | @expr.App _ _ _ s2 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s2 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x * x0)%expr
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | (_ @ _)%expr_pat =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if
- args0 =?
- 2 ^ (2 * Z.log2_up args0 / 2) - 1
- then
- match
- invert_low
- (2 * Z.log2_up args0) args1
- with
- | Some y =>
- UnderLets.Base
+ (type.base t4) ->
+ defaults.expr
+ (type.base base.type.Z))
+ =>
+ match a0 with
+ | Some x'0 =>
+ if
+ (args0 =?
+ 2 ^ (2 * Z.log2_up args0 / 2) -
+ 1) &&
+ (args2 =?
+ 2 ^ (2 * Z.log2_up args0 / 2) -
+ 1)
+ then
+ Some
+ (UnderLets.Base
(#(ident.fancy_mulll
(2 * Z.log2_up args0))%expr @
- (x' v, ##(y)%expr))%expr_pat
- | None =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v0,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value'
- false s2 ->
- Compile.value' true
- d2 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- end
- else
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up args0 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v0,
- ##(y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s2 _
- (@expr.App _ _ _ s3 _ #(idc1)%expr_pat x5) x4 =>
- match
- match idc1 with
- | ident.Z_land => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
+ (x' v, x'0 v0))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun
+ _ : value' false s3 -> value' true d3
+ => None
+ end (reflect x4)
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun _ : value' false s3 -> value' true d3 => None
+ end (reflect x2)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
+ | _ => None
+ end
+ | _ => None
+ end;;
+ match x2 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match x0 with
+ | (@expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t1 idc1) x4 @ x3)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc1 pident.Z_shiftr;
+ match x3 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pattern.ident.LiteralZ;
+ match
+ s as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t3 =>
+ fun v : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr (type.base base.type.Z))
+ =>
+ match a with
+ | Some x' =>
+ match
+ s2 as t4
+ return
+ (value' false t4 ->
+ option
+ (UnderLets.UnderLets base.type ident
+ var
+ (defaults.expr
+ (type.base base.type.Z))))
+ with
+ | type.base t4 =>
+ fun v0 : defaults.expr (type.base t4) =>
+ base.try_make_transport_cps t4
+ base.type.Z
+ (fun
+ a0 : option
(defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s2 as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (x' v,
- x'0 v0))%expr_pat
- else
- match x4 with
- | #(idc3)%expr_pat =>
- match
- match
- idc3
- with
- | @ident.Literal
- t6 v1 =>
- match
- t6 as t7
- return
- (base.base_interp
- t7 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v2 : Z =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s3 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'1 v1,
- x'2 v2))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 ->
- d4)%ptype =>
- fun
- _ :
- Compile.value'
- false s4 ->
- Compile.value'
- true d4
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x5)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 ->
- d4)%ptype =>
- fun
- _ :
- Compile.value'
- false s4 ->
- Compile.value'
- true d4
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value'
- true d4 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- match x4 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s3 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ :
- Compile.value'
- false s4 ->
- Compile.value'
- true d4
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x5)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value' true
- d4 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | ($_)%expr =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s3 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v1 : defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base t6))
- t5 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v0,
- x'0 v1))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value'
- true d4 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s3 as t4
- return
- (Compile.value'
- false t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun
- t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value'
- true d4 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s3 as t4
- return
- (Compile.value'
- false t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun
- t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s5 -> d5)%ptype =>
- fun
- _ : Compile.value'
- false s5 ->
- Compile.value'
- true d5 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s5 -> d5)%ptype =>
- fun
- _ : Compile.value' false s5 ->
- Compile.value' true d5
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s3 as t4
- return
- (Compile.value'
- false t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun
- t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^
- (2 * args2 /
- 2) - 1
- then
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- args2))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value'
- true d4 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- end
- | @expr.App _ _ _ s2 _
- (@expr.App _ _ _ s3 _ ($_)%expr _) _ | @expr.App _
- _ _ s2 _
- (@expr.App _ _ _ s3 _ (@expr.Abs _ _ _ _ _ _) _)
- _ | @expr.App _ _ _ s2 _
- (@expr.App _ _ _ s3 _ (_ @ _)%expr_pat _) _ |
- @expr.App _ _ _ s2 _
- (@expr.App _ _ _ s3 _ (@expr.LetIn _ _ _ _ _ _ _)
- _) _ => UnderLets.Base (x * x0)%expr
- | @expr.App _ _ _ s2 _ #(_)%expr_pat _ | @expr.App _
- _ _ s2 _ ($_)%expr _ | @expr.App _ _ _ s2 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s2 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x * x0)%expr
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if
- args0 =?
- 2 ^ (2 * Z.log2_up args0 / 2) - 1
- then
- match
- invert_low
- (2 * Z.log2_up args0) args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 * Z.log2_up args0))%expr @
- (x' v, ##(y)%expr))%expr_pat
- | None =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1
- then
- match
- invert_high
- (2 *
- Z.log2_up
- args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v0,
- ##
- (y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- end
- else
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^ (2 *
- Z.log2_up args0 /
- 2) - 1
- then
- match
- invert_high
- (2 *
- Z.log2_up args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mullh
- (2 *
- Z.log2_up
- args0))%expr @
- (x'0 v0,
- ##(y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ #(idc1)%expr_pat x5) x4 =>
- match
- match idc1 with
- | ident.Z_land => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
+ (type.base t4) ->
+ defaults.expr
+ (type.base base.type.Z))
+ =>
+ match a0 with
+ | Some x'0 =>
+ if
+ args0 =? 2 ^ (2 * args2 / 2) - 1
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_mullh
+ (2 * args2))%expr @
+ (x' v, x'0 v0))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun
+ _ : value' false s3 -> value' true d3
+ => None
+ end (reflect x4)
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun _ : value' false s3 -> value' true d3 => None
+ end (reflect x1)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
+ | _ => None
+ end
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match x0 with
+ | (@expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t1 idc1) x4 @ x3)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc1 pident.Z_shiftr;
+ match x3 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pattern.ident.LiteralZ;
+ match
+ s0 as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t3 =>
+ fun v : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr (type.base base.type.Z))
+ =>
+ match a with
+ | Some x' =>
+ match
+ s2 as t4
+ return
+ (value' false t4 ->
+ option
+ (UnderLets.UnderLets base.type ident
+ var
+ (defaults.expr
+ (type.base base.type.Z))))
+ with
+ | type.base t4 =>
+ fun v0 : defaults.expr (type.base t4) =>
+ base.try_make_transport_cps t4
+ base.type.Z
+ (fun
+ a0 : option
(defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s1 as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^ (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (x' v,
- x'0 v0))%expr_pat
- else
- match x4 with
- | #(idc3)%expr_pat =>
- match
- match
- idc3
- with
- | @ident.Literal
- t6 v1 =>
- match
- t6 as t7
- return
- (base.base_interp
- t7 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v2 : Z =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s2 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args3 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x'1 v1,
- x'2 v2))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x5)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- match x4 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x5)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | ($_)%expr =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v1 : defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base t6))
- t5 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v0,
- x'0 v1))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t4
- return
- (Compile.value'
- false t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun
- t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t4
- return
- (Compile.value'
- false t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun
- t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- (args0 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
- then
- UnderLets.Base
- (#
- (ident.fancy_mulll
- (2 *
- Z.log2_up
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value'
- true d4 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t4
- return
- (Compile.value'
- false t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun
- t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- 2
- ^
- (2 * args2 /
- 2) - 1
- then
- UnderLets.Base
- (#
- (ident.fancy_mullh
- (2 *
- args2))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x5)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- end
- | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ ($_)%expr _) _ | @expr.App _
- _ _ s1 _
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _)
- _ | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ (_ @ _)%expr_pat _) _ |
- @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _)
- _) _ => UnderLets.Base (x * x0)%expr
- | @expr.App _ _ _ s1 _ #(_)%expr_pat _ | @expr.App _
- _ _ s1 _ ($_)%expr _ | @expr.App _ _ _ s1 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s1 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x * x0)%expr
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | None =>
- match
- match idc with
- | ident.Z_shiftr => Some tt
+ (type.base t4) ->
+ defaults.expr
+ (type.base base.type.Z))
+ =>
+ match a0 with
+ | Some x'0 =>
+ if
+ args0 =? 2 ^ (2 * args2 / 2) - 1
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_mullh
+ (2 * args2))%expr @
+ (x' v, x'0 v0))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun
+ _ : value' false s3 -> value' true d3
+ => None
+ end (reflect x4)
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun _ : value' false s3 -> value' true d3 => None
+ end (reflect x2)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
| _ => None
end
- with
- | Some _ =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
+ | _ => None
+ end);;
+ _ <- pattern.ident.invert_bind_args idc pident.Z_shiftr;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ (y <- invert_low (2 * args0) args1;
+ Some
+ (Some
+ (UnderLets.Base
+ (#(ident.fancy_mulhl (2 * args0))%expr @
+ (x' v, (##y)%expr))%expr_pat)));;;
+ None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2);;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ (y <- invert_high (2 * args0) args1;
+ Some
+ (Some
+ (UnderLets.Base
+ (#(ident.fancy_mulhh (2 * args0))%expr @
+ (x' v, (##y)%expr))%expr_pat)));;;
+ None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2)
+ | @expr.App _ _ _ s1 _
+ (@expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t1 idc1) x4) x3 =>
+ (_ <- pattern.ident.invert_bind_args idc1 pident.Z_land;
+ match x4 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pattern.ident.LiteralZ;
+ match
+ s0 as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t3 =>
+ fun v : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr (type.base base.type.Z))
+ =>
+ match a with
+ | Some x' =>
+ match
+ s1 as t4
+ return
+ (value' false t4 ->
+ option
+ (UnderLets.UnderLets base.type ident
+ var
+ (defaults.expr
+ (type.base base.type.Z))))
+ with
+ | type.base t4 =>
+ fun v0 : defaults.expr (type.base t4) =>
+ base.try_make_transport_cps t4
+ base.type.Z
+ (fun
+ a0 : option
+ (defaults.expr
+ (type.base t4) ->
+ defaults.expr
+ (type.base base.type.Z))
+ =>
+ match a0 with
+ | Some x'0 =>
+ if
+ args2 =? 2 ^ (2 * args0 / 2) - 1
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_mulhl
+ (2 * args0))%expr @
+ (x' v, x'0 v0))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun
+ _ : value' false s3 -> value' true d3
+ => None
+ end (reflect x3)
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun _ : value' false s3 -> value' true d3 => None
+ end (reflect x2)
+ | _ => None
+ end;;
+ match x3 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pattern.ident.LiteralZ;
+ match
+ s0 as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t3 =>
+ fun v : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr (type.base base.type.Z))
+ =>
+ match a with
+ | Some x' =>
+ match
+ s2 as t4
+ return
+ (value' false t4 ->
+ option
+ (UnderLets.UnderLets base.type ident
+ var
+ (defaults.expr
+ (type.base base.type.Z))))
+ with
+ | type.base t4 =>
+ fun v0 : defaults.expr (type.base t4) =>
+ base.try_make_transport_cps t4
+ base.type.Z
+ (fun
+ a0 : option
+ (defaults.expr
+ (type.base t4) ->
+ defaults.expr
+ (type.base base.type.Z))
+ =>
+ match a0 with
+ | Some x'0 =>
+ if
+ args2 =? 2 ^ (2 * args0 / 2) - 1
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_mulhl
+ (2 * args0))%expr @
+ (x' v, x'0 v0))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun
+ _ : value' false s3 -> value' true d3
+ => None
+ end (reflect x4)
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun _ : value' false s3 -> value' true d3 => None
+ end (reflect x2)
+ | _ => None
+ end);;
+ _ <- pattern.ident.invert_bind_args idc1 pident.Z_shiftr;
+ match x3 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pattern.ident.LiteralZ;
+ match
+ s0 as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t3 =>
+ fun v : defaults.expr (type.base t3) =>
+ base.try_make_transport_cps t3 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr (type.base base.type.Z))
+ =>
+ match a with
+ | Some x' =>
+ match
+ s2 as t4
return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
+ (value' false t4 ->
+ option
+ (UnderLets.UnderLets base.type ident
+ var
+ (defaults.expr
+ (type.base base.type.Z))))
+ with
+ | type.base t4 =>
+ fun v0 : defaults.expr (type.base t4) =>
+ base.try_make_transport_cps t4
base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
(fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- match
- invert_low (2 * args0) args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mulhl
+ a0 : option
+ (defaults.expr (type.base t4) ->
+ defaults.expr
+ (type.base base.type.Z))
+ =>
+ match a0 with
+ | Some x'0 =>
+ if args0 =? args2
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_mulhh
(2 * args0))%expr @
- (x' v, ##(y)%expr))%expr_pat
- | None =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- match
- invert_high
- (2 * args0)
- args1
- with
- | Some y =>
- UnderLets.Base
- (#(ident.fancy_mulhh
- (2 *
- args0))%expr @
- (x'0 v0,
- ##(y)%expr))%expr_pat
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false
- s1 ->
- Compile.value' true
- d1 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- end
- | None => UnderLets.Base (x * x0)%expr
+ (x' v, x'0 v0))%expr_pat)
+ else None
+ | None => None
end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None => UnderLets.Base (x * x0)%expr
- end
- | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ #(idc1)%expr_pat x4) x3 =>
- match
- match idc1 with
- | ident.Z_land => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- match
- s1 as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args2 =?
- 2
- ^ (2 * args0 /
- 2) - 1
- then
- UnderLets.Base
- (#(ident.fancy_mulhl
- (2 *
- args0))%expr @
- (x' v,
- x'0 v0))%expr_pat
- else
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match
- idc3
- with
- | @ident.Literal
- t6 v1 =>
- match
- t6 as t7
- return
- (base.base_interp
- t7 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v2 : Z =>
- Some v2
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v1
- | _ =>
- None
- end
- with
- | Some
- args3 =>
- match
- s0 as t6
- return
- (Compile.value'
- false t6 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t6 =>
- fun
- v1 :
- defaults.expr
- (type.base
- t6) =>
- base.try_make_transport_cps
- (fun
- t7 : base.type
- =>
- defaults.expr
- (type.base
- t7)) t6
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a1 :
- option
- (defaults.expr
- (type.base
- t6) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a1
- with
- | Some
- x'1 =>
- match
- s2 as t7
- return
- (Compile.value'
- false t7 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base
- t7 =>
- fun
- v2 :
- defaults.expr
- (type.base
- t7) =>
- base.try_make_transport_cps
- (fun
- t8 : base.type
- =>
- defaults.expr
- (type.base
- t8)) t7
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a2 :
- option
- (defaults.expr
- (type.base
- t7) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a2
- with
- | Some
- x'2 =>
- if
- args3 =?
- 2
- ^
- (2 *
- args0 / 2) -
- 1
- then
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- args0))%expr @
- (
- x'1 v1,
- x'2 v2))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x2)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end
- | _ =>
- UnderLets.Base
- (x * x0)%expr
- end
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- args2 =?
- 2
- ^
- (2 *
- args0 / 2) -
- 1
- then
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect
- x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | ($_)%expr =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t5 =>
- fun
- v1 : defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base t6))
- t5 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args2 =?
- 2
- ^
- (2 * args0 /
- 2) - 1
- then
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- args0))%expr @
- (
- x' v0,
- x'0 v1))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t4
- return
- (Compile.value'
- false t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun
- t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args2 =?
- 2
- ^
- (2 * args0 /
- 2) - 1
- then
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t4
- return
- (Compile.value'
- false t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun
- t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args2 =?
- 2
- ^
- (2 * args0 /
- 2) - 1
- then
- UnderLets.Base
- (#
- (ident.fancy_mulhl
- (2 *
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value'
- true d4 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- match
- s2 as t4
- return
- (Compile.value'
- false t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- base.type.Z)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base
- t4) =>
- base.try_make_transport_cps
- (fun
- t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args0 =?
- args2
- then
- UnderLets.Base
- (#
- (ident.fancy_mulhh
- (2 *
- args0))%expr @
- (
- x' v,
- x'0 v0))%expr_pat
- else
- UnderLets.Base
- (x * x0)%expr
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (x * x0)%expr
- end
- (Compile.reflect x4)
- | None =>
- UnderLets.Base
- (x * x0)%expr
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base (x * x0)%expr
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- end
- | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ ($_)%expr _) _ | @expr.App _
- _ _ s1 _
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _)
- _ | @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ (_ @ _)%expr_pat _) _ |
- @expr.App _ _ _ s1 _
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _)
- _) _ => UnderLets.Base (x * x0)%expr
- | @expr.App _ _ _ s1 _ #(_)%expr_pat _ | @expr.App _
- _ _ s1 _ ($_)%expr _ | @expr.App _ _ _ s1 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s1 _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (x * x0)%expr
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- end
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App _
- _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ | @expr.App
- _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ | @expr.App _ _ _
- s _ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ =>
- UnderLets.Base (x * x0)%expr
- | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr _ |
- @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => UnderLets.Base (x * x0)%expr
- | _ => UnderLets.Base (x * x0)%expr
- end
+ | (s3 -> d3)%ptype =>
+ fun _ : value' false s3 -> value' true d3
+ => None
+ end (reflect x4)
+ | None => None
+ end)
+ | (s3 -> d3)%ptype =>
+ fun _ : value' false s3 -> value' true d3 => None
+ end (reflect x2)
+ | _ => None
+ end
+ | @expr.App _ _ _ s1 _ (@expr.App _ _ _ s2 _ ($_)%expr _) _ |
+ @expr.App _ _ _ s1 _
+ (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _) _ | @expr.App
+ _ _ _ s1 _ (@expr.App _ _ _ s2 _ (_ @ _)%expr_pat _) _ |
+ @expr.App _ _ _ s1 _
+ (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ => None
+ | @expr.App _ _ _ s1 _ #(_)%expr_pat _ | @expr.App _ _ _ s1 _
+ ($_)%expr _ | @expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ |
+ @expr.App _ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end
+ | _ => None
+ end
+ | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App _
+ _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ | @expr.App
+ _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ | @expr.App _ _
+ _ s _ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ => None
+ | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr
+ _ | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s
+ _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;;
+ UnderLets.Base (x * x0)%expr)%option
| ident.Z_pow =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat
@@ -17745,17081 +1535,750 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
| ident.Z_add_get_carry =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ match x1 with
+ | (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x3 @ x2)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.Z_shiftl;
+ match x2 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_add (Z.log2 args) args1)%expr @
+ (x0, x' v))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x3)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
| _ => None
- end
- with
- | Some args =>
- match x1 with
- | #(_)%expr_pat =>
- match x0 with
- | (@expr.App _ _ _ s0 _ #(idc1) x3 @ x2)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) args1)%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | ($_)%expr =>
- match x0 with
- | (@expr.App _ _ _ s0 _ #(idc0) x3 @ x2)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) args1)%expr @
- (x1, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x1, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match x0 with
- | (@expr.App _ _ _ s1 _ #(idc0) x3 @ x2)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) args1)%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s1 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | (#(_) @ _)%expr_pat =>
- match x0 with
- | (@expr.App _ _ _ s1 _ #(idc1) x4 @ x3)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) args1)%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s1 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | (($_)%expr @ _)%expr_pat =>
- match x0 with
- | (@expr.App _ _ _ s1 _ #(idc0) x4 @ x3)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) args1)%expr @
- (x1, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x1, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s1 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat =>
- match x0 with
- | (@expr.App _ _ _ s2 _ #(idc0) x4 @ x3)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) args1)%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | (@expr.App _ _ _ s0 _ #(idc0) x3 @ x2)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args)
- args1)%expr @ (x0, x' v))%expr_pat
- else
- match x0 with
- | (@expr.App _ _ _ s2 _ #(idc2) x5 @
- x4)%expr_pat =>
- match
- match idc2 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t5
- v0 =>
- match
- t5 as t6
- return
- (base.base_interp
- t6 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit =>
- None
- | base.type.Z =>
- fun v1 : Z =>
- Some v1
- | base.type.bool =>
- fun _ : bool =>
- None
- | base.type.nat =>
- fun _ : nat =>
- None
- end v0
- | _ => None
- end
- with
- | Some args3 =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t5 =>
- fun
- v0 : defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_add
- (Z.log2
- args)
- args3)%expr @
- (x1,
- x'0 v0))%expr_pat
- else
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#
- (ident.Z_add_get_carry)%expr @
- x @ x0 @
- x1)%expr_pat
- | None =>
- UnderLets.Base
- (#
- (ident.Z_add_get_carry)%expr @
- x @ x0 @
- x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- (Compile.reflect x5)
- | None =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc2 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t5
- v0 =>
- match
- t5 as t6
- return
- (base.base_interp
- t6 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun v1 : Z
- =>
- Some v1
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun _ : nat
- => None
- end v0
- | _ => None
- end
- with
- | Some args3 =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t5 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_add
- (Z.log2
- args)
- (- args3))%expr @
- (x1,
- x'0 v0))%expr_pat
- else
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#
- (ident.Z_add_get_carry)%expr @
- x @ x0 @
- x1)%expr_pat
- | None =>
- UnderLets.Base
- (#
- (ident.Z_add_get_carry)%expr @
- x @ x0 @
- x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @
- x1)%expr_pat
- end
- (Compile.reflect
- x5)
- | None =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | None =>
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @
- _)%expr_pat |
- (@expr.App _ _ _ s2 _
- (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _
- (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- match x0 with
- | (@expr.App _ _ _ s2 _ #(idc2) x5 @ x4)%expr_pat =>
- match
- match idc2 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- args2)%expr @
- (x1, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc2 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args)
- (- args2))%expr @
- (x1, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _)
- _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _
- (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- end
- | ($_)%expr =>
- match x0 with
- | (@expr.App _ _ _ s2 _ #(idc1) x5 @ x4)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- args2)%expr @
- (x1, x' v0))%expr_pat
- else
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args2))%expr @
- (x1, x' v0))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _)
- _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match x0 with
- | (@expr.App _ _ _ s3 _ #(idc1) x5 @ x4)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s3 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- args2)%expr @
- (x1, x' v))%expr_pat
- else
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s3 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args2))%expr @
- (x1, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s3 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.LetIn _ _ _ _ _ _ _)
- _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | (_ @ _)%expr_pat =>
- match x0 with
- | (@expr.App _ _ _ s3 _ #(idc1) x6 @ x5)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s3 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- args2)%expr @
- (x1, x' v))%expr_pat
- else
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s3 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args2))%expr @
- (x1, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s3 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.LetIn _ _ _ _ _ _ _)
- _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x0 with
- | (@expr.App _ _ _ s2 _ #(idc1) x6 @ x5)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- args2)%expr @
- (x1, x' v))%expr_pat
- else
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args2))%expr @
- (x1, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _)
- _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x0 with
- | #(_)%expr_pat =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x0, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | ($_)%expr =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x0, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | (#(_) @ _)%expr_pat =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x0, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | (($_)%expr @ _)%expr_pat =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x0, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x0, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | (@expr.App _ _ _ s2 _ #(idc1) x5 @ x4)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- args2)%expr @
- (x1, x' v))%expr_pat
- else
- match x2 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t5
- v0 =>
- match
- t5 as t6
- return
- (base.base_interp
- t6 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun v1 : Z
- =>
- Some v1
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun _ : nat
- => None
- end v0
- | _ => None
- end
- with
- | Some args3 =>
- match
- s0 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t5 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_add
- (Z.log2
- args)
- (- args3))%expr @
- (x0,
- x'0 v0))%expr_pat
- else
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#
- (ident.Z_add_get_carry)%expr @
- x @ x0 @
- x1)%expr_pat
- | None =>
- UnderLets.Base
- (#
- (ident.Z_add_get_carry)%expr @
- x @ x0 @
- x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @
- x1)%expr_pat
- end
- (Compile.reflect
- x3)
- | None =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- match x2 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args)
- (- args2))%expr @
- (x0, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- end
- | ($_)%expr =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args2))%expr @
- (x0, x' v0))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args2))%expr @
- (x0, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args2))%expr @
- (x0, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args2))%expr @
- (x0, x' v))%expr_pat
- else
- match x4 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal
- t5 v0 =>
- match
- t5 as t6
- return
- (base.base_interp
- t6 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v1 : Z =>
- Some v1
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v0
- | _ => None
- end
- with
- | Some args3 =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base
- t5 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_add
- (Z.log2
- args)
- (- args3))%expr @
- (x1,
- x'0 v0))%expr_pat
- else
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#
- (ident.Z_add_get_carry)%expr @
- x @ x0 @
- x1)%expr_pat
- | None =>
- UnderLets.Base
- (#
- (ident.Z_add_get_carry)%expr @
- x @ x0 @
- x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (#
- (ident.Z_add_get_carry)%expr @
- x @ x0 @
- x1)%expr_pat
- end
- (Compile.reflect
- x5)
- | None =>
- if
- args =?
- 2
- ^ Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#
- (ident.Z_add_get_carry)%expr @
- x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- match x4 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp
- t5 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit =>
- None
- | base.type.Z =>
- fun v0 : Z =>
- Some v0
- | base.type.bool =>
- fun _ : bool =>
- None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2
- ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args)
- (- args2))%expr @
- (x1, x' v))%expr_pat
- else
- if
- args =?
- 2
- ^ Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#
- (ident.Z_add_get_carry)%expr @
- x @ x0 @
- x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- end
- | ($_)%expr =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args)
- (- args2))%expr @
- (x1, x' v0))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args)
- (- args2))%expr @
- (x1, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args)
- (- args2))%expr @
- (x1, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2
- args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value' true
- d4 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- end
- | None =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x0, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x0, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x0, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _)
- _ @ _)%expr_pat =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x0, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x0, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x0, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- end
- | None =>
- match x0 with
- | (@expr.App _ _ _ s2 _ #(idc1) x5 @ x4)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- args1)%expr @
- (x1, x' v))%expr_pat
- else
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x1, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _)
- _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- end
- | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat =>
- match x0 with
- | (@expr.App _ _ _ s2 _ #(idc0) x5 @ x4)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) args1)%expr @
- (x1, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x1, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat =>
- match x0 with
- | (@expr.App _ _ _ s3 _ #(idc0) x5 @ x4)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s3 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) args1)%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s3 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s3 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat =>
- match x0 with
- | (@expr.App _ _ _ s3 _ #(idc0) x6 @ x5)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s3 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) args1)%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s3 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s3 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- match x0 with
- | (@expr.App _ _ _ s2 _ #(idc0) x6 @ x5)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) args1)%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | (@expr.LetIn _ _ _ _ _ _ _ @ _)%expr_pat =>
- match x0 with
- | (@expr.App _ _ _ s1 _ #(idc0) x5 @ x4)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) args1)%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s1 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x0 with
- | (@expr.App _ _ _ s0 _ #(idc0) x4 @ x3)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) args1)%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args)
- (- args1))%expr @
- (x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @
- x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
+ end;;
+ match x0 with
+ | (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x3 @ x2)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.Z_shiftl;
+ match x2 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_add (Z.log2 args) args1)%expr @
+ (x1, x' v))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x3)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
+ | _ => None
+ end;;
+ match x1 with
+ | (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x3 @ x2)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.Z_shiftr;
+ match x2 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_add (Z.log2 args) (- args1))%expr @
+ (x0, x' v))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x3)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
+ | _ => None
+ end;;
+ match x0 with
+ | (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x3 @ x2)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.Z_shiftr;
+ match x2 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_add (Z.log2 args) (- args1))%expr @
+ (x1, x' v))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x3)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
+ | _ => None
+ end;;
+ (if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_add (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat)
+ else None)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat)%option
| ident.Z_add_with_carry =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
UnderLets.Base (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
| ident.Z_add_with_get_carry =>
fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ match x2 with
+ | (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x4 @ x3)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.Z_shiftl;
+ match x3 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_addc (Z.log2 args) args1)%expr @
+ (x0, x1, x' v))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x4)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
| _ => None
- end
- with
- | Some args =>
- match x2 with
- | #(_)%expr_pat =>
- match x1 with
- | (@expr.App _ _ _ s0 _ #(idc1) x4 @ x3)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) args1)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | ($_)%expr =>
- match x1 with
- | (@expr.App _ _ _ s0 _ #(idc0) x4 @ x3)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) args1)%expr @
- (x0, x2, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x2, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match x1 with
- | (@expr.App _ _ _ s1 _ #(idc0) x4 @ x3)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) args1)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s1 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | (#(_) @ _)%expr_pat =>
- match x1 with
- | (@expr.App _ _ _ s1 _ #(idc1) x5 @ x4)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) args1)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s1 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | (($_)%expr @ _)%expr_pat =>
- match x1 with
- | (@expr.App _ _ _ s1 _ #(idc0) x5 @ x4)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) args1)%expr @
- (x0, x2, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x2, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s1 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat =>
- match x1 with
- | (@expr.App _ _ _ s2 _ #(idc0) x5 @ x4)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) args1)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | (@expr.App _ _ _ s0 _ #(idc0) x4 @ x3)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- args1)%expr @ (x0, x1, x' v))%expr_pat
- else
- match x1 with
- | (@expr.App _ _ _ s2 _ #(idc2) x6 @
- x5)%expr_pat =>
- match
- match idc2 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t5
- v0 =>
- match
- t5 as t6
- return
- (base.base_interp
- t6 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit =>
- None
- | base.type.Z =>
- fun v1 : Z =>
- Some v1
- | base.type.bool =>
- fun _ : bool =>
- None
- | base.type.nat =>
- fun _ : nat =>
- None
- end v0
- | _ => None
- end
- with
- | Some args3 =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t5 =>
- fun
- v0 : defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_addc
- (Z.log2
- args)
- args3)%expr @
- (x0, x2,
- x'0 v0))%expr_pat
- else
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1,
- x2))%expr_pat
- else
- UnderLets.Base
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- end
- (Compile.reflect x6)
- | None =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc2 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t5
- v0 =>
- match
- t5 as t6
- return
- (base.base_interp
- t6 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun v1 : Z
- =>
- Some v1
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun _ : nat
- => None
- end v0
- | _ => None
- end
- with
- | Some args3 =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t5 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_addc
- (Z.log2
- args)
- (- args3))%expr @
- (x0, x2,
- x'0 v0))%expr_pat
- else
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1,
- x2))%expr_pat
- else
- UnderLets.Base
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- end
- (Compile.reflect
- x6)
- | None =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- end
- | _ =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @
- _)%expr_pat |
- (@expr.App _ _ _ s2 _
- (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _
- (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- match x1 with
- | (@expr.App _ _ _ s2 _ #(idc2) x6 @ x5)%expr_pat =>
- match
- match idc2 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- args2)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc2 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args)
- (- args2))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _)
- _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _
- (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- end
- | ($_)%expr =>
- match x1 with
- | (@expr.App _ _ _ s2 _ #(idc1) x6 @ x5)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- args2)%expr @
- (x0, x2, x' v0))%expr_pat
- else
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args2))%expr @
- (x0, x2, x' v0))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _)
- _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | @expr.Abs _ _ _ _ _ _ =>
- match x1 with
- | (@expr.App _ _ _ s3 _ #(idc1) x6 @ x5)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s3 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- args2)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s3 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args2))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s3 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.LetIn _ _ _ _ _ _ _)
- _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | (_ @ _)%expr_pat =>
- match x1 with
- | (@expr.App _ _ _ s3 _ #(idc1) x7 @ x6)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x6 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s3 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- args2)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x7)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x6 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s3 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args2))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x7)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s3 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.LetIn _ _ _ _ _ _ _)
- _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x1 with
- | (@expr.App _ _ _ s2 _ #(idc1) x7 @ x6)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x6 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- args2)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x7)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x6 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args2))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x7)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _)
- _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x1 with
- | #(_)%expr_pat =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | ($_)%expr =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x1, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | (#(_) @ _)%expr_pat =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | (($_)%expr @ _)%expr_pat =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x1, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | (@expr.App _ _ _ s2 _ #(idc1) x6 @ x5)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- args2)%expr @
- (x0, x2, x' v))%expr_pat
- else
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t5
- v0 =>
- match
- t5 as t6
- return
- (base.base_interp
- t6 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun v1 : Z
- =>
- Some v1
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun _ : nat
- => None
- end v0
- | _ => None
- end
- with
- | Some args3 =>
- match
- s0 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t5 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_addc
- (Z.log2
- args)
- (- args3))%expr @
- (x0, x1,
- x'0 v0))%expr_pat
- else
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1,
- x2))%expr_pat
- else
- UnderLets.Base
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- end
- (Compile.reflect
- x4)
- | None =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- end
- | _ =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- match x3 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args)
- (- args2))%expr @
- (x0, x1, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- end
- | ($_)%expr =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4) =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args2))%expr @
- (x0, x1, x' v0))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args2))%expr @
- (x0, x1, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args2))%expr @
- (x0, x1, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args2))%expr @
- (x0, x1, x' v))%expr_pat
- else
- match x5 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal
- t5 v0 =>
- match
- t5 as t6
- return
- (base.base_interp
- t6 ->
- option Z)
- with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.Z =>
- fun
- v1 : Z =>
- Some v1
- | base.type.bool =>
- fun
- _ : bool
- => None
- | base.type.nat =>
- fun
- _ : nat
- => None
- end v0
- | _ => None
- end
- with
- | Some args3 =>
- match
- s2 as t5
- return
- (Compile.value'
- false t5 ->
- UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base
- t5 =>
- fun
- v0 :
- defaults.expr
- (type.base
- t5) =>
- base.try_make_transport_cps
- (fun
- t6 : base.type
- =>
- defaults.expr
- (type.base
- t6)) t5
- base.type.Z
- (UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a0 :
- option
- (defaults.expr
- (type.base
- t5) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match
- a0
- with
- | Some
- x'0 =>
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_addc
- (Z.log2
- args)
- (- args3))%expr @
- (x0, x2,
- x'0 v0))%expr_pat
- else
- if
- args =?
- 2
- ^
- Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1,
- x2))%expr_pat
- else
- UnderLets.Base
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- UnderLets.Base
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- end
- (Compile.reflect
- x6)
- | None =>
- if
- args =?
- 2
- ^ Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1,
- x2))%expr_pat
- else
- UnderLets.Base
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- end
- | _ =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- match x5 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp
- t5 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit =>
- None
- | base.type.Z =>
- fun v0 : Z =>
- Some v0
- | base.type.bool =>
- fun _ : bool =>
- None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t4
- return
- (Compile.value' false
- t4 ->
- UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type
- =>
- defaults.expr
- (type.base t5))
- t4 base.type.Z
- (UnderLets.UnderLets
- base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2
- ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args)
- (- args2))%expr @
- (x0, x2,
- x' v))%expr_pat
- else
- if
- args =?
- 2
- ^ Z.log2
- args
- then
- UnderLets.Base
- (#
- (ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1,
- x2))%expr_pat
- else
- UnderLets.Base
- (#
- (ident.Z_add_with_get_carry)%expr @
- x @ x0 @
- x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- end
- | ($_)%expr =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t4 v0 =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t4
- return
- (Compile.value' false t4 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t4 =>
- fun
- v0 : defaults.expr
- (type.base t4)
- =>
- base.try_make_transport_cps
- (fun t5 : base.type =>
- defaults.expr
- (type.base t5)) t4
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t4) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args)
- (- args2))%expr @
- (x0, x2,
- x' v0))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args)
- (- args2))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value' true
- d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args)
- (- args2))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2
- args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @
- x2)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value'
- false s4 ->
- Compile.value' true
- d4 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- end
- | None =>
- match x3 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x1, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x1, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _)
- _ @ _)%expr_pat =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- end
- | None =>
- match x1 with
- | (@expr.App _ _ _ s2 _ #(idc1) x6 @ x5)%expr_pat =>
- match
- match idc1 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4))
- t3 base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- args1)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if
- args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc1 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun
- v : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z *
- base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a with
- | Some x' =>
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if
- args =?
- 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3
- =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _)
- _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- end
- | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat =>
- match x1 with
- | (@expr.App _ _ _ s2 _ #(idc0) x6 @ x5)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return (base.base_interp t4 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) args1)%expr @
- (x0, x2, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3)
- =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x2, x' v0))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat =>
- match x1 with
- | (@expr.App _ _ _ s3 _ #(idc0) x6 @ x5)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s3 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) args1)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s3 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s3 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat =>
- match x1 with
- | (@expr.App _ _ _ s3 _ #(idc0) x7 @ x6)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x6 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s3 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) args1)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x7)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x6 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s3 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s4 -> d4)%ptype =>
- fun
- _ : Compile.value' false s4 ->
- Compile.value' true d4 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x7)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s3 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s3 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- match x1 with
- | (@expr.App _ _ _ s2 _ #(idc0) x7 @ x6)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x6 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) args1)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x7)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x6 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s2 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value' false s3 ->
- Compile.value' true d3 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x7)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s2 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | (@expr.LetIn _ _ _ _ _ _ _ @ _)%expr_pat =>
- match x1 with
- | (@expr.App _ _ _ s1 _ #(idc0) x6 @ x5)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) args1)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x5 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s1 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x6)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s1 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x1 with
- | (@expr.App _ _ _ s0 _ #(idc0) x5 @ x4)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) args1)%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x4 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return
- (base.base_interp t3 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
- =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base t2) ->
- defaults.expr
- (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args)
- (- args1))%expr @
- (x0, x2, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x5)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
+ end;;
+ match x1 with
+ | (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x4 @ x3)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.Z_shiftl;
+ match x3 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_addc (Z.log2 args) args1)%expr @
+ (x0, x2, x' v))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x4)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
+ | _ => None
+ end;;
+ match x2 with
+ | (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x4 @ x3)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.Z_shiftr;
+ match x3 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_addc (Z.log2 args)
+ (- args1))%expr @ (x0, x1, x' v))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x4)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
+ | _ => None
+ end;;
+ match x1 with
+ | (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x4 @ x3)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.Z_shiftr;
+ match x3 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_addc (Z.log2 args)
+ (- args1))%expr @ (x0, x2, x' v))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x4)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
+ | _ => None
+ end;;
+ (if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_addc (Z.log2 args) 0)%expr @ (x0, x1, x2))%expr_pat)
+ else None)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat)%option
| ident.Z_sub_get_borrow =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ match x1 with
+ | (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x3 @ x2)%expr_pat =>
+ (_ <- pattern.ident.invert_bind_args idc0 pident.Z_shiftl;
+ match x2 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_sub (Z.log2 args) args1)%expr @
+ (x0, x' v))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x3)
+ | _ => None
+ end);;
+ _ <- pattern.ident.invert_bind_args idc0 pident.Z_shiftr;
+ match x2 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_sub (Z.log2 args) (- args1))%expr @
+ (x0, x' v))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x3)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
| _ => None
- end
- with
- | Some args =>
- match x1 with
- | (@expr.App _ _ _ s0 _ #(idc0) x3 @ x2)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_sub (Z.log2 args)
- args1)%expr @ (x0, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_sub (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @
- x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_sub (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_sub (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_sub
- (Z.log2 args) (- args1))%expr @
- (x0, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_sub
- (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @
- x @ x0 @ x1)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @
- x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_sub (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
- x1)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_sub (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_sub (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- end
- end
- | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_sub (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_sub (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- end
+ end;;
+ (if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_sub (Z.log2 args) 0)%expr @ (x0, x1))%expr_pat)
+ else None)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat)%option
| ident.Z_sub_with_get_borrow =>
fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ match x2 with
+ | (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x4 @ x3)%expr_pat =>
+ (_ <- pattern.ident.invert_bind_args idc0 pident.Z_shiftl;
+ match x3 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_subb (Z.log2 args) args1)%expr @
+ (x0, x1, x' v))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x4)
+ | _ => None
+ end);;
+ _ <- pattern.ident.invert_bind_args idc0 pident.Z_shiftr;
+ match x3 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.Z * base.type.Z)%etype))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_subb (Z.log2 args)
+ (- args1))%expr @ (x0, x1, x' v))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x4)
+ | _ => None
+ end
+ | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
+ None
| _ => None
- end
- with
- | Some args =>
- match x2 with
- | (@expr.App _ _ _ s0 _ #(idc0) x4 @ x3)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftl => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_subb (Z.log2 args)
- args1)%expr @ (x0, x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_subb (Z.log2 args)
- 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_subb (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_subb (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- end
- | None =>
- match
- match idc0 with
- | ident.Z_shiftr => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base
- (base.type.Z * base.type.Z)%etype)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_subb
- (Z.log2 args) (- args1))%expr @
- (x0, x1, x' v))%expr_pat
- else
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_subb
- (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end (Compile.reflect x4)
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_subb (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_subb (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | None =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_subb (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- end
- end
- | (@expr.App _ _ _ s0 _ ($_)%expr _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (_ @ _) _ @ _)%expr_pat |
- (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_subb (Z.log2 args) 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | _ =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_subb (Z.log2 args) 0)%expr @ (x0, x1, x2))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
+ end;;
+ (if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_subb (Z.log2 args) 0)%expr @ (x0, x1, x2))%expr_pat)
+ else None)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat)%option
| ident.Z_zselect =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- match x with
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x3) x2 =>
- match match idc with
- | ident.Z_cc_m => Some tt
- | _ => None
- end with
- | Some _ =>
+ (((match x with
+ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc) x3) x2 =>
+ (_ <- pattern.ident.invert_bind_args idc pident.Z_cc_m;
match x3 with
- | #(idc0)%expr_pat =>
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
with
- | Some args0 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type => defaults.expr (type.base t1))
- t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args0 =? 2 ^ Z.log2 args0
- then
- UnderLets.Base
- (#(ident.fancy_selm (Z.log2 args0))%expr @
- (x' v, x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- end
- | None =>
- match
- match idc with
- | ident.Z_land => Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x3 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2 base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base base.type.Z))
- =>
- match a with
- | Some x' =>
- if args0 =? 1
- then
- UnderLets.Base
- (#(ident.fancy_sell)%expr @
- (x' v, x0, x1))%expr_pat
- else
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t3 v0 =>
- match
- t3 as t4
- return
- (base.base_interp t4 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v1 : Z => Some v1
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base base.type.Z)))
- with
- | type.base t3 =>
- fun
- v0 : defaults.expr
- (type.base t3) =>
- base.try_make_transport_cps
- (fun t4 : base.type =>
- defaults.expr
- (type.base t4)) t3
- base.type.Z
- (UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base
- base.type.Z)))
- (fun
- a0 : option
- (defaults.expr
- (type.base t3) ->
- defaults.expr
- (type.base
- base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args1 =? 1
- then
- UnderLets.Base
- (#(ident.fancy_sell)%expr @
- (x'0 v0, x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.fancy_selc)%expr @
- (x, x0, x1))%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @
- x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1
- =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @
- x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @
- (x, x0, x1))%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @
- (x, x0, x1))%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x2)
- | None =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident
- var
- (defaults.expr
- (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args0 =? 1
- then
- UnderLets.Base
- (#(ident.fancy_sell)%expr @
- (x' v, x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.fancy_selc)%expr @
- (x, x0, x1))%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @
- x0 @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- end
- end
- | ($_)%expr =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t2 v0 =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v1 : Z => Some v1
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v0
- | _ => None
- end
- with
- | Some args0 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v0 : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t3 : base.type =>
- defaults.expr (type.base t3)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args0 =? 1
- then
- UnderLets.Base
- (#(ident.fancy_sell)%expr @
- (x' v0, x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.fancy_selc)%expr @
- (x, x0, x1))%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @
- x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- end
- | @expr.LetIn _ _ _ _ _ _ _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2
- return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args0 =? 1
- then
- UnderLets.Base
- (#(ident.fancy_sell)%expr @
- (x' v, x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.fancy_selc)%expr @
- (x, x0, x1))%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @
- x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- end
- | _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2
- return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- base.try_make_transport_cps
- (fun t1 : base.type =>
- defaults.expr (type.base t1)) t2
- base.type.Z
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base base.type.Z)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr
- (type.base base.type.Z)) =>
- match a with
- | Some x' =>
- if args0 =? 1
- then
- UnderLets.Base
- (#(ident.fancy_sell)%expr @
- (x' v, x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.fancy_selc)%expr @
- (x, x0, x1))%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @
- x1)%expr_pat
- end)
- | (s2 -> d2)%ptype =>
- fun
- _ : Compile.value' false s2 ->
- Compile.value' true d2 =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- end
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- end
- end
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App _
- _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ | @expr.App
- _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ | @expr.App _ _ _
- s _ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ =>
- UnderLets.Base (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr _ |
- @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- | _ => UnderLets.Base (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
- end
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 =? 2 ^ Z.log2 args0
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_selm (Z.log2 args0))%expr @
+ (x' v, x0, x1))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2)
+ | _ => None
+ end);;
+ _ <- pattern.ident.invert_bind_args idc pident.Z_land;
+ match x3 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match
+ s as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 =? 1
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_sell)%expr @ (x' v, x0, x1))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2)
+ | _ => None
+ end;;
+ match x2 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ base.try_make_transport_cps t2 base.type.Z
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base base.type.Z)) =>
+ match a with
+ | Some x' =>
+ if args0 =? 1
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_sell)%expr @ (x' v, x0, x1))%expr_pat)
+ else None
+ | None => None
+ end)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x3)
+ | _ => None
+ end
+ | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
+ _ _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ |
+ @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ |
+ @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ => None
+ | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr
+ _ | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _
+ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ Some
+ (UnderLets.Base (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat));;
+ None);;;
+ UnderLets.Base (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat)%option
| ident.Z_add_modulo =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
UnderLets.Base (#(ident.fancy_addm)%expr @ (x, x0, x1))%expr_pat
| ident.Z_rshi =>
fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_rshi (Z.log2 args) args0)%expr @
- (x0, x1))%expr_pat
- else
- UnderLets.Base
- (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- | None =>
- UnderLets.Base
- (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- UnderLets.Base (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
+ (match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralZ;
+ match x2 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0
+ pattern.ident.LiteralZ;
+ (if args =? 2 ^ Z.log2 args
+ then
+ Some
+ (UnderLets.Base
+ (#(ident.fancy_rshi (Z.log2 args) args0)%expr @ (x0, x1))%expr_pat)
+ else None)
+ | _ => None
+ end
+ | _ => None
+ end;;;
+ UnderLets.Base (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat)%option
| ident.Z_cc_m =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat
@@ -34881,4 +2340,4 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat
end
- : Compile.value' true t
+ : value' true t
diff --git a/src/Experiments/NewPipeline/nbe_rewrite_head.out b/src/Experiments/NewPipeline/nbe_rewrite_head.out
index 82d4983bd..e5808c478 100644
--- a/src/Experiments/NewPipeline/nbe_rewrite_head.out
+++ b/src/Experiments/NewPipeline/nbe_rewrite_head.out
@@ -1,5 +1,5 @@
nbe_rewrite_head =
-match idc in (ident t) return (Compile.value' true t) with
+match idc in (ident t) return (value' true t) with
| @ident.Literal t v =>
match
t as t0
@@ -8,225 +8,91 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base t0)))
with
- | base.type.unit => fun v0 : unit => UnderLets.Base ##(v0)%expr
- | base.type.Z => fun v0 : Z => UnderLets.Base ##(v0)%expr
- | base.type.bool => fun v0 : bool => UnderLets.Base ##(v0)%expr
- | base.type.nat => fun v0 : nat => UnderLets.Base ##(v0)%expr
+ | base.type.unit => fun v0 : unit => UnderLets.Base (##v0)%expr
+ | base.type.Z => fun v0 : Z => UnderLets.Base (##v0)%expr
+ | base.type.bool => fun v0 : bool => UnderLets.Base (##v0)%expr
+ | base.type.nat => fun v0 : nat => UnderLets.Base (##v0)%expr
end v
| ident.Nat_succ =>
fun x : defaults.expr (type.base base.type.nat) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args => UnderLets.Base ##(Nat.succ args)%expr
- | None => UnderLets.Base (#(ident.Nat_succ)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Nat_succ)%expr @ x)%expr_pat
- end
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralNat;
+ Some (UnderLets.Base (##(Nat.succ args))%expr)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Nat_succ)%expr @ x)%expr_pat)%option
| ident.Nat_pred =>
fun x : defaults.expr (type.base base.type.nat) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args => UnderLets.Base ##(Nat.pred args)%expr
- | None => UnderLets.Base (#(ident.Nat_pred)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Nat_pred)%expr @ x)%expr_pat
- end
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralNat;
+ Some (UnderLets.Base (##(Nat.pred args))%expr)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Nat_pred)%expr @ x)%expr_pat)%option
| ident.Nat_max =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralNat;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralNat;
+ Some (UnderLets.Base (##(Nat.max args args0))%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option nat)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##(Nat.max args args0)%expr
- | None =>
- UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat
- end
- | None => UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat)%option
| ident.Nat_mul =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralNat;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralNat;
+ Some (UnderLets.Base (##(args * args0)%nat)%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option nat)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##((args * args0)%nat)%expr
- | None =>
- UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat
- end
- | None => UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat)%option
| ident.Nat_add =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralNat;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralNat;
+ Some (UnderLets.Base (##(args + args0)%nat)%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option nat)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##((args + args0)%nat)%expr
- | None =>
- UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat
- end
- | None => UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat)%option
| ident.Nat_sub =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralNat;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralNat;
+ Some (UnderLets.Base (##(args - args0)%nat)%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option nat)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##((args - args0)%nat)%expr
- | None =>
- UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat
- end
- | None => UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat)%option
| @ident.nil t => UnderLets.Base []%expr_pat
| @ident.cons t =>
fun (x : defaults.expr (type.base t))
@@ -237,132 +103,120 @@ match idc in (ident t) return (Compile.value' true t) with
=> UnderLets.Base (x, x0)%expr_pat
| @ident.fst A B =>
fun x : defaults.expr (type.base (A * B)%etype) =>
- match x with
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x1) x0 =>
- match
- match idc with
- | @ident.pair A0 B0 => Some (A0, B0)
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base A)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base A)))
- with
- | type.base t3 =>
- fun _ : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t2
- A
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base A)))
- (fun
- a : option
- (defaults.expr (type.base t2) ->
- defaults.expr (type.base A)) =>
- match a with
- | Some x' => UnderLets.Base (x' v)
- | None =>
- UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1
- => UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- end (Compile.reflect x0)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1 =>
- UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- end (Compile.reflect x1)
- | None => UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- end
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App _
- _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ | @expr.App
- _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ | @expr.App _ _ _
- s _ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ =>
- UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr _ |
- @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- | _ => UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
- end
+ ((match x with
+ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc) x1) x0 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.pair;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base A))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ match
+ s as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base A))))
+ with
+ | type.base t3 =>
+ fun _ : defaults.expr (type.base t3) =>
+ Some
+ (base.try_make_transport_cps t2 A
+ (fun
+ a : option
+ (defaults.expr (type.base t2) ->
+ defaults.expr (type.base A)) =>
+ match a with
+ | Some x' => UnderLets.Base (x' v)
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.fst) @ (??, ??))
+ (#(ident.fst)%expr @ x)%expr_pat)
+ end))
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x0)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
+ _ _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ |
+ @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ |
+ @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ => None
+ | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr
+ _ | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s
+ _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat)%option
| @ident.snd A B =>
fun x : defaults.expr (type.base (A * B)%etype) =>
- match x with
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x1) x0 =>
- match
- match idc with
- | @ident.pair A0 B0 => Some (A0, B0)
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base B)))
- with
- | type.base t2 =>
- fun _ : defaults.expr (type.base t2) =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base B)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0)) t3
- B
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base B)))
- (fun
- a : option
- (defaults.expr (type.base t3) ->
- defaults.expr (type.base B)) =>
- match a with
- | Some x' => UnderLets.Base (x' v0)
- | None =>
- UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1
- => UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- end (Compile.reflect x0)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1 =>
- UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- end (Compile.reflect x1)
- | None => UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- end
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App _
- _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ | @expr.App
- _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ | @expr.App _ _ _
- s _ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ =>
- UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr _ |
- @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- | _ => UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
- end
+ ((match x with
+ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc) x1) x0 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.pair;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base B))))
+ with
+ | type.base t2 =>
+ fun _ : defaults.expr (type.base t2) =>
+ match
+ s as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base B))))
+ with
+ | type.base t3 =>
+ fun v0 : defaults.expr (type.base t3) =>
+ Some
+ (base.try_make_transport_cps t3 B
+ (fun
+ a : option
+ (defaults.expr (type.base t3) ->
+ defaults.expr (type.base B)) =>
+ match a with
+ | Some x' => UnderLets.Base (x' v0)
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.snd) @ (??, ??))
+ (#(ident.snd)%expr @ x)%expr_pat)
+ end))
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x0)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x1)
+ | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
+ _ _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ |
+ @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ |
+ @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ => None
+ | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr
+ _ | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s
+ _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat)%option
| @ident.prod_rect A B T =>
fun
(x : defaults.expr (type.base A) ->
@@ -370,63 +224,53 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base T)))
(x0 : defaults.expr (type.base (A * B)%etype)) =>
- match x0 with
- | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr _ |
- @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ =>
- UnderLets.Base
- (#(ident.prod_rect)%expr @
- (λ (x2 : var (type.base A))(x3 : var (type.base B)),
- UnderLets.to_expr (x ($x2) ($x3)))%expr @ x0)%expr_pat
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x2) x1 =>
- match
- match idc with
- | @ident.pair A0 B0 => Some (A0, B0)
- | _ => None
- end
- with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base T)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base T)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- Compile.castbe v
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base T)))
- (fun x3 : option (defaults.expr (type.base A)) =>
- match x3 with
- | Some x4 =>
- Compile.castbe v0
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base T)))
- (fun y : option (defaults.expr (type.base B)) =>
- match y with
- | Some y0 =>
+ ((match x0 with
+ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc) x2) x1 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.pair;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base T))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ match
+ s as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base T))))
+ with
+ | type.base t3 =>
+ fun v0 : defaults.expr (type.base t3) =>
+ castbe v
+ (option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base T))))
+ (fun x3 : option (defaults.expr (type.base A)) =>
+ (x4 <- x3;
+ Some
+ (castbe v0
+ (option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base T))))
+ (fun y : option (defaults.expr (type.base B)) =>
+ (y0 <- y;
+ Some
+ (Some
(fv <-- (e <-- x x4 y0;
UnderLets.Base
{|
anyexpr_ty := T;
unwrap := e |});
base.try_make_transport_cps
- (fun t0 : base.type =>
- defaults.expr (type.base t0))
(let (anyexpr_ty, _) := fv in anyexpr_ty)
T
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base T)))
(fun
a : option
(defaults.expr
@@ -453,79 +297,39 @@ match idc in (ident t) return (Compile.value' true t) with
unwrap))
| None =>
UnderLets.Base
- (#(ident.prod_rect)%expr @
- (λ (x5 : var (type.base A))
- (x6 : var (type.base B)),
- UnderLets.to_expr
- (x ($x5) ($x6)))%expr @ x0)%expr_pat
- end))%under_lets
- | None =>
- UnderLets.Base
- (#(ident.prod_rect)%expr @
- (λ (x5 : var (type.base A))(x6 :
- var
- (type.base
- B)),
- UnderLets.to_expr (x ($x5) ($x6)))%expr @
- x0)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.prod_rect)%expr @
- (λ (x4 : var (type.base A))(x5 : var
- (type.base B)),
- UnderLets.to_expr (x ($x4) ($x5)))%expr @ x0)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1
- =>
- UnderLets.Base
- (#(ident.prod_rect)%expr @
- (λ (x3 : var (type.base A))(x4 : var (type.base B)),
- UnderLets.to_expr (x ($x3) ($x4)))%expr @ x0)%expr_pat
- end (Compile.reflect x1)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.prod_rect)%expr @
- (λ (x3 : var (type.base A))(x4 : var (type.base B)),
- UnderLets.to_expr (x ($x3) ($x4)))%expr @ x0)%expr_pat
- end (Compile.reflect x2)
- | None =>
- UnderLets.Base
- (#(ident.prod_rect)%expr @
- (λ (x3 : var (type.base A))(x4 : var (type.base B)),
- UnderLets.to_expr (x ($x3) ($x4)))%expr @ x0)%expr_pat
- end
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App _
- _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ =>
- UnderLets.Base
- (#(ident.prod_rect)%expr @
- (λ (x3 : var (type.base A))(x4 : var (type.base B)),
- UnderLets.to_expr (x ($x3) ($x4)))%expr @ x0)%expr_pat
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ |
- @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ =>
- UnderLets.Base
- (#(ident.prod_rect)%expr @
- (λ (x4 : var (type.base A))(x5 : var (type.base B)),
- UnderLets.to_expr (x ($x4) ($x5)))%expr @ x0)%expr_pat
- | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base
- (#(ident.prod_rect)%expr @
- (λ (x3 : var (type.base A))(x4 : var (type.base B)),
- UnderLets.to_expr (x ($x3) ($x4)))%expr @ x0)%expr_pat
- | @expr.LetIn _ _ _ _ _ _ _ =>
- UnderLets.Base
- (#(ident.prod_rect)%expr @
- (λ (x2 : var (type.base A))(x3 : var (type.base B)),
- UnderLets.to_expr (x ($x2) ($x3)))%expr @ x0)%expr_pat
- | _ =>
- UnderLets.Base
- (#(ident.prod_rect)%expr @
- (λ (x1 : var (type.base A))(x2 : var (type.base B)),
- UnderLets.to_expr (x ($x1) ($x2)))%expr @ x0)%expr_pat
- end
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.prod_rect) @
+ ??{?? -> ?? -> ??} @
+ (??, ??))
+ (#(ident.prod_rect)%expr @
+ (λ (x5 : var (type.base A))
+ (x6 : var (type.base B)),
+ UnderLets.to_expr
+ (x ($x5) ($x6)))%expr @ x0)%expr_pat)
+ end))%under_lets));;;
+ None)));;;
+ None)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x1)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2)
+ | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
+ _ _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ |
+ @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ |
+ @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ => None
+ | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr
+ _ | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s
+ _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base
+ (#(ident.prod_rect)%expr @
+ (λ (x1 : var (type.base A))(x2 : var (type.base B)),
+ UnderLets.to_expr (x ($x1) ($x2)))%expr @ x0)%expr_pat)%option
| @ident.bool_rect T =>
fun
(x
@@ -533,35 +337,22 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base T)))
(x1 : defaults.expr (type.base base.type.bool)) =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option bool) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun v0 : bool => Some v0
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
+ ((match x1 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc
+ pattern.ident.LiteralBool;
+ Some
(fv <-- (e <-- (if args as b
return
(UnderLets.UnderLets base.type ident var
(defaults.expr
(type.base (if b then T else T))))
- then x ##(tt)%expr
- else x0 ##(tt)%expr);
+ then x (##tt)%expr
+ else x0 (##tt)%expr);
UnderLets.Base
{| anyexpr_ty := if args then T else T; unwrap := e |});
base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
(let (anyexpr_ty, _) := fv in anyexpr_ty) T
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base T)))
(fun
a : option
(defaults.expr
@@ -582,35 +373,24 @@ match idc in (ident t) return (Compile.value' true t) with
unwrap))
| None =>
UnderLets.Base
- (#(ident.bool_rect)%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x2)))%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.bool_rect) @ ??{() -> ??} @ ??{() -> ??} @
+ #(pattern.ident.LiteralBool))
+ (#(ident.bool_rect)%expr @
+ (λ x2 : var (type.base base.type.unit),
+ UnderLets.to_expr (x ($x2)))%expr @
+ (λ x2 : var (type.base base.type.unit),
+ UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat)
end))%under_lets
- | None =>
- UnderLets.Base
- (#(ident.bool_rect)%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x2)))%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat
- end
- | ($_)%expr | @expr.Abs _ _ _ _ _ _ =>
- UnderLets.Base
- (#(ident.bool_rect)%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x2)))%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat
- | _ =>
- UnderLets.Base
- (#(ident.bool_rect)%expr @
- (λ x3 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x3)))%expr @
- (λ x3 : var (type.base base.type.unit),
- UnderLets.to_expr (x0 ($x3)))%expr @ x1)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base
+ (#(ident.bool_rect)%expr @
+ (λ x2 : var (type.base base.type.unit),
+ UnderLets.to_expr (x ($x2)))%expr @
+ (λ x2 : var (type.base base.type.unit),
+ UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat)%option
| @ident.nat_rect P =>
fun
(x : defaults.expr (type.base base.type.unit) ->
@@ -621,49 +401,36 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base P)))
(x1 : defaults.expr (type.base base.type.nat)) =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args =>
- Compile.castv x0
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base P)))
- (fun
- S_case : option
- (defaults.expr (type.base base.type.nat) ->
- defaults.expr (type.base P) ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base P))) =>
- match S_case with
- | Some S_case0 =>
+ ((match x1 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralNat;
+ castv x0
+ (option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base P))))
+ (fun
+ S_case : option
+ (defaults.expr (type.base base.type.nat) ->
+ defaults.expr (type.base P) ->
+ UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base P))) =>
+ (S_case0 <- S_case;
+ Some
+ (Some
(fv <-- (e <-- nat_rect
(fun _ : nat =>
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base P)))
- (x ##(tt)%expr)
+ (x (##tt)%expr)
(fun (n' : nat)
(rec : UnderLets.UnderLets base.type
ident var
(defaults.expr (type.base P)))
=> rec0 <-- rec;
- S_case0 ##(n')%expr rec0) args;
+ S_case0 (##n')%expr rec0) args;
UnderLets.Base {| anyexpr_ty := P; unwrap := e |});
base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
(let (anyexpr_ty, _) := fv in anyexpr_ty) P
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base P)))
(fun
a : option
(defaults.expr
@@ -684,46 +451,28 @@ match idc in (ident t) return (Compile.value' true t) with
unwrap))
| None =>
UnderLets.Base
- (#(ident.nat_rect)%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x2)))%expr @
- (λ (x2 : var (type.base base.type.nat))
- (x3 : var (type.base P)),
- UnderLets.to_expr (x0 ($x2) ($x3)))%expr @ x1)%expr_pat
- end))%under_lets
- | None =>
- UnderLets.Base
- (#(ident.nat_rect)%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x2)))%expr @
- (λ (x2 : var (type.base base.type.nat))(x3 : var
- (type.base
- P)),
- UnderLets.to_expr (x0 ($x2) ($x3)))%expr @ x1)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.nat_rect)%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x2)))%expr @
- (λ (x2 : var (type.base base.type.nat))(x3 : var (type.base P)),
- UnderLets.to_expr (x0 ($x2) ($x3)))%expr @ x1)%expr_pat
- end
- | ($_)%expr | @expr.Abs _ _ _ _ _ _ =>
- UnderLets.Base
- (#(ident.nat_rect)%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x2)))%expr @
- (λ (x2 : var (type.base base.type.nat))(x3 : var (type.base P)),
- UnderLets.to_expr (x0 ($x2) ($x3)))%expr @ x1)%expr_pat
- | _ =>
- UnderLets.Base
- (#(ident.nat_rect)%expr @
- (λ x3 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x3)))%expr @
- (λ (x3 : var (type.base base.type.nat))(x4 : var (type.base P)),
- UnderLets.to_expr (x0 ($x3) ($x4)))%expr @ x1)%expr_pat
- end
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.nat_rect) @ ??{() -> ??} @
+ ??{type.base base.type.nat -> ?? -> ??} @
+ #(pattern.ident.LiteralNat))
+ (#(ident.nat_rect)%expr @
+ (λ x2 : var (type.base base.type.unit),
+ UnderLets.to_expr (x ($x2)))%expr @
+ (λ (x2 : var (type.base base.type.nat))
+ (x3 : var (type.base P)),
+ UnderLets.to_expr (x0 ($x2) ($x3)))%expr @
+ x1)%expr_pat)
+ end))%under_lets));;;
+ None)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base
+ (#(ident.nat_rect)%expr @
+ (λ x2 : var (type.base base.type.unit),
+ UnderLets.to_expr (x ($x2)))%expr @
+ (λ (x2 : var (type.base base.type.nat))(x3 : var (type.base P)),
+ UnderLets.to_expr (x0 ($x2) ($x3)))%expr @ x1)%expr_pat)%option
| @ident.nat_rect_arrow P Q =>
fun
(x : defaults.expr (type.base P) ->
@@ -738,182 +487,114 @@ match idc in (ident t) return (Compile.value' true t) with
(defaults.expr (type.base Q)))
(x1 : defaults.expr (type.base base.type.nat))
(x2 : defaults.expr (type.base P)) =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args =>
- Compile.castv x0
+ (match x1 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralNat;
+ castv x0
+ (option
(UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base Q)))
- (fun
- S_case : option
- (defaults.expr (type.base base.type.nat) ->
- (defaults.expr (type.base P) ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base Q))) ->
- defaults.expr (type.base P) ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base Q))) =>
- match S_case with
- | Some S_case0 =>
- Compile.castbe x2
+ (defaults.expr (type.base Q))))
+ (fun
+ S_case : option
+ (defaults.expr (type.base base.type.nat) ->
+ (defaults.expr (type.base P) ->
+ UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base Q))) ->
+ defaults.expr (type.base P) ->
+ UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base Q))) =>
+ (S_case0 <- S_case;
+ Some
+ (castbe x2
+ (option
(UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base Q)))
- (fun v : option (defaults.expr (type.base P)) =>
- match v with
- | Some v0 =>
- (fv <-- (e <-- nat_rect
- (fun _ : nat =>
- defaults.expr (type.base P) ->
- UnderLets.UnderLets base.type
- ident var
- (defaults.expr (type.base Q)))
- x
- (fun (n' : nat)
- (rec : defaults.expr
- (type.base P) ->
- UnderLets.UnderLets
- base.type ident var
- (defaults.expr
- (type.base Q)))
- (v1 : defaults.expr
- (type.base P)) =>
- S_case0 ##(n')%expr rec v1) args
- v0;
- UnderLets.Base
- {| anyexpr_ty := Q; unwrap := e |});
- base.try_make_transport_cps
- (fun t0 : base.type =>
- defaults.expr (type.base t0))
- (let (anyexpr_ty, _) := fv in anyexpr_ty) Q
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base Q)))
- (fun
- a : option
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := fv in
- anyexpr_ty)) ->
- defaults.expr (type.base Q)) =>
- match a with
- | Some x' =>
- UnderLets.Base
- (x'
- (let
- (anyexpr_ty, unwrap) as a0
- return
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) :=
- a0 in
- anyexpr_ty))) := fv in
- unwrap))
- | None =>
+ (defaults.expr (type.base Q))))
+ (fun v : option (defaults.expr (type.base P)) =>
+ (v0 <- v;
+ Some
+ (Some
+ (fv <-- (e <-- nat_rect
+ (fun _ : nat =>
+ defaults.expr (type.base P) ->
+ UnderLets.UnderLets base.type
+ ident var
+ (defaults.expr (type.base Q))) x
+ (fun (n' : nat)
+ (rec : defaults.expr
+ (type.base P) ->
+ UnderLets.UnderLets
+ base.type ident var
+ (defaults.expr
+ (type.base Q)))
+ (v1 : defaults.expr
+ (type.base P)) =>
+ S_case0 (##n')%expr rec v1) args
+ v0;
UnderLets.Base
- (#(ident.nat_rect_arrow)%expr @
- (λ x3 : var (type.base P),
- UnderLets.to_expr (x ($x3)))%expr @
- (λ (x3 : var (type.base base.type.nat))
- (x4 : var
- (type.base P -> type.base Q)%ptype)
- (x5 : var (type.base P)),
- UnderLets.to_expr
- (x0 ($x3)
- (fun
- x6 : defaults.expr
- (type.base P) =>
- UnderLets.Base
- ($x4 @ x6)%expr_pat) ($x5)))%expr @
- x1 @ x2)%expr_pat
- end))%under_lets
- | None =>
- UnderLets.Base
- (#(ident.nat_rect_arrow)%expr @
- (λ x3 : var (type.base P),
- UnderLets.to_expr (x ($x3)))%expr @
- (λ (x3 : var (type.base base.type.nat))(x4 :
- var
- (type.base
- P ->
- type.base
- Q)%ptype)
- (x5 : var (type.base P)),
- UnderLets.to_expr
- (x0 ($x3)
- (fun x6 : defaults.expr (type.base P) =>
- UnderLets.Base ($x4 @ x6)%expr_pat)
- ($x5)))%expr @ x1 @ x2)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.nat_rect_arrow)%expr @
- (λ x3 : var (type.base P),
- UnderLets.to_expr (x ($x3)))%expr @
- (λ (x3 : var (type.base base.type.nat))(x4 : var
- (type.base
- P ->
- type.base
- Q)%ptype)
- (x5 : var (type.base P)),
- UnderLets.to_expr
- (x0 ($x3)
- (fun x6 : defaults.expr (type.base P) =>
- UnderLets.Base ($x4 @ x6)%expr_pat) ($x5)))%expr @
- x1 @ x2)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.nat_rect_arrow)%expr @
- (λ x3 : var (type.base P),
- UnderLets.to_expr (x ($x3)))%expr @
- (λ (x3 : var (type.base base.type.nat))(x4 : var
- (type.base P ->
- type.base Q)%ptype)
- (x5 : var (type.base P)),
- UnderLets.to_expr
- (x0 ($x3)
- (fun x6 : defaults.expr (type.base P) =>
- UnderLets.Base ($x4 @ x6)%expr_pat) ($x5)))%expr @ x1 @
- x2)%expr_pat
- end
- | ($_)%expr | @expr.Abs _ _ _ _ _ _ =>
- UnderLets.Base
- (#(ident.nat_rect_arrow)%expr @
- (λ x3 : var (type.base P),
- UnderLets.to_expr (x ($x3)))%expr @
- (λ (x3 : var (type.base base.type.nat))(x4 : var
- (type.base P ->
- type.base Q)%ptype)
- (x5 : var (type.base P)),
- UnderLets.to_expr
- (x0 ($x3)
- (fun x6 : defaults.expr (type.base P) =>
- UnderLets.Base ($x4 @ x6)%expr_pat) ($x5)))%expr @ x1 @ x2)%expr_pat
- | _ =>
- UnderLets.Base
- (#(ident.nat_rect_arrow)%expr @
- (λ x4 : var (type.base P),
- UnderLets.to_expr (x ($x4)))%expr @
- (λ (x4 : var (type.base base.type.nat))(x5 : var
- (type.base P ->
- type.base Q)%ptype)
- (x6 : var (type.base P)),
- UnderLets.to_expr
- (x0 ($x4)
- (fun x7 : defaults.expr (type.base P) =>
- UnderLets.Base ($x5 @ x7)%expr_pat) ($x6)))%expr @ x1 @ x2)%expr_pat
- end
+ {| anyexpr_ty := Q; unwrap := e |});
+ base.try_make_transport_cps
+ (let (anyexpr_ty, _) := fv in anyexpr_ty) Q
+ (fun
+ a : option
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) := fv in
+ anyexpr_ty)) ->
+ defaults.expr (type.base Q)) =>
+ match a with
+ | Some x' =>
+ UnderLets.Base
+ (x'
+ (let
+ (anyexpr_ty, unwrap) as a0
+ return
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) :=
+ a0 in
+ anyexpr_ty))) := fv in
+ unwrap))
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.nat_rect_arrow) @
+ ??{?? -> ??} @
+ ??{type.base base.type.nat ->
+ (?? -> ??) -> ?? -> ??} @
+ #(pattern.ident.LiteralNat) @
+ ??)
+ (#(ident.nat_rect_arrow)%expr @
+ (λ x3 : var (type.base P),
+ UnderLets.to_expr (x ($x3)))%expr @
+ (λ (x3 : var (type.base base.type.nat))
+ (x4 : var
+ (type.base P -> type.base Q)%ptype)
+ (x5 : var (type.base P)),
+ UnderLets.to_expr
+ (x0 ($x3)
+ (fun
+ x6 : defaults.expr
+ (type.base P) =>
+ UnderLets.Base
+ ($x4 @ x6)%expr_pat)
+ ($x5)))%expr @ x1 @ x2)%expr_pat)
+ end))%under_lets));;;
+ None)));;;
+ None)
+ | _ => None
+ end;;;
+ UnderLets.Base
+ (#(ident.nat_rect_arrow)%expr @
+ (λ x3 : var (type.base P),
+ UnderLets.to_expr (x ($x3)))%expr @
+ (λ (x3 : var (type.base base.type.nat))(x4 : var
+ (type.base P ->
+ type.base Q)%ptype)
+ (x5 : var (type.base P)),
+ UnderLets.to_expr
+ (x0 ($x3)
+ (fun x6 : defaults.expr (type.base P) =>
+ UnderLets.Base ($x4 @ x6)%expr_pat) ($x5)))%expr @ x1 @ x2)%expr_pat)%option
| @ident.list_rect A P =>
fun
(x : defaults.expr (type.base base.type.unit) ->
@@ -925,96 +606,94 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base P)))
(x1 : defaults.expr (type.base (base.type.list A))) =>
- Compile.castv x0
- (UnderLets.UnderLets base.type ident var (defaults.expr (type.base P)))
- (fun
- Pcons : option
- (defaults.expr (type.base A) ->
- defaults.expr (type.base (base.type.list A)) ->
- defaults.expr (type.base P) ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base P))) =>
- match Pcons with
- | Some Pcons0 =>
- reflect_list_cps x1
- (fun ls : option (list (defaults.expr (type.base A))) =>
- match ls with
- | Some ls0 =>
- (fv <-- (e <-- list_rect
- (fun
- _ : list (defaults.expr (type.base A))
- =>
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base P)))
- (x ##(tt)%expr)
- (fun (x2 : defaults.expr (type.base A))
- (xs : list
- (defaults.expr (type.base A)))
- (rec : UnderLets.UnderLets base.type
- ident var
- (defaults.expr (type.base P)))
- =>
- rec' <-- rec;
- Pcons0 x2 (reify_list xs) rec') ls0;
- UnderLets.Base {| anyexpr_ty := P; unwrap := e |});
- base.try_make_transport_cps
- (fun t : base.type => defaults.expr (type.base t))
- (let (anyexpr_ty, _) := fv in anyexpr_ty) P
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base P)))
- (fun
- a : option
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := fv in anyexpr_ty)) ->
- defaults.expr (type.base P)) =>
- match a with
- | Some x' =>
- UnderLets.Base
- (x'
- (let
- (anyexpr_ty, unwrap) as a0
- return
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := a0 in
- anyexpr_ty))) := fv in
- unwrap))
- | None =>
- UnderLets.Base
- (#(ident.list_rect)%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x2)))%expr @
- (λ (x2 : var (type.base A))(x3 : var
- (type.base
- (base.type.list
+ ((castv x0
+ (option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base P))))
+ (fun
+ Pcons : option
+ (defaults.expr (type.base A) ->
+ defaults.expr (type.base (base.type.list A)) ->
+ defaults.expr (type.base P) ->
+ UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base P))) =>
+ (Pcons0 <- Pcons;
+ Some
+ (reflect_list_cps x1
+ (fun ls : option (list (defaults.expr (type.base A))) =>
+ (ls0 <- ls;
+ Some
+ (Some
+ (fv <-- (e <-- list_rect
+ (fun
+ _ : list
+ (defaults.expr (type.base A))
+ =>
+ UnderLets.UnderLets base.type ident
+ var (defaults.expr (type.base P)))
+ (x (##tt)%expr)
+ (fun
+ (x2 : defaults.expr (type.base A))
+ (xs : list
+ (defaults.expr
+ (type.base A)))
+ (rec : UnderLets.UnderLets
+ base.type ident var
+ (defaults.expr
+ (type.base P))) =>
+ rec' <-- rec;
+ Pcons0 x2 (reify_list xs) rec') ls0;
+ UnderLets.Base
+ {| anyexpr_ty := P; unwrap := e |});
+ base.try_make_transport_cps
+ (let (anyexpr_ty, _) := fv in anyexpr_ty) P
+ (fun
+ a : option
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) := fv in
+ anyexpr_ty)) ->
+ defaults.expr (type.base P)) =>
+ match a with
+ | Some x' =>
+ UnderLets.Base
+ (x'
+ (let
+ (anyexpr_ty, unwrap) as a0
+ return
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) := a0 in
+ anyexpr_ty))) := fv in
+ unwrap))
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.list_rect) @ ??{() -> ??} @
+ ??{?? -> ?? -> ?? -> ??} @
+ ??{type.base (pattern.base.type.list ??)})
+ (#(ident.list_rect)%expr @
+ (λ x2 : var (type.base base.type.unit),
+ UnderLets.to_expr (x ($x2)))%expr @
+ (λ (x2 : var (type.base A))(x3 :
+ var
+ (type.base
+ (base.type.list
A)))
- (x4 : var (type.base P)),
- UnderLets.to_expr (x0 ($x2) ($x3) ($x4)))%expr @
- x1)%expr_pat
- end))%under_lets
- | None =>
- UnderLets.Base
- (#(ident.list_rect)%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x2)))%expr @
- (λ (x2 : var (type.base A))(x3 : var
- (type.base
- (base.type.list A)))
- (x4 : var (type.base P)),
- UnderLets.to_expr (x0 ($x2) ($x3) ($x4)))%expr @ x1)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.list_rect)%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x2)))%expr @
- (λ (x2 : var (type.base A))(x3 : var
- (type.base
- (base.type.list A)))
- (x4 : var (type.base P)),
- UnderLets.to_expr (x0 ($x2) ($x3) ($x4)))%expr @ x1)%expr_pat
- end)
+ (x4 : var (type.base P)),
+ UnderLets.to_expr (x0 ($x2) ($x3) ($x4)))%expr @
+ x1)%expr_pat)
+ end))%under_lets));;;
+ None)));;;
+ None);;
+ None);;;
+ UnderLets.Base
+ (#(ident.list_rect)%expr @
+ (λ x2 : var (type.base base.type.unit),
+ UnderLets.to_expr (x ($x2)))%expr @
+ (λ (x2 : var (type.base A))(x3 : var (type.base (base.type.list A)))
+ (x4 : var (type.base P)),
+ UnderLets.to_expr (x0 ($x2) ($x3) ($x4)))%expr @ x1)%expr_pat)%option
| @ident.list_case A P =>
fun
(x : defaults.expr (type.base base.type.unit) ->
@@ -1025,20 +704,14 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base P)))
(x1 : defaults.expr (type.base (base.type.list A))) =>
- match x1 with
- | #(idc)%expr_pat =>
- match match idc with
- | @ident.nil t0 => Some t0
- | _ => None
- end with
- | Some _ =>
- (fv <-- (e <-- x ##(tt)%expr;
+ ((match x1 with
+ | @expr.Ident _ _ _ t idc =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.nil;
+ Some
+ (fv <-- (e <-- x (##tt)%expr;
UnderLets.Base {| anyexpr_ty := P; unwrap := e |});
base.try_make_transport_cps
- (fun t0 : base.type => defaults.expr (type.base t0))
(let (anyexpr_ty, _) := fv in anyexpr_ty) P
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base P)))
(fun
a : option
(defaults.expr
@@ -1059,85 +732,67 @@ match idc in (ident t) return (Compile.value' true t) with
unwrap))
| None =>
UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x2)))%expr @
- (λ (x2 : var (type.base A))(x3 : var
- (type.base
- (base.type.list
- A))),
- UnderLets.to_expr (x0 ($x2) ($x3)))%expr @ x1)%expr_pat
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.list_case) @ ??{() -> ??} @
+ ??{?? -> ?? -> ??} @ [])
+ (#(ident.list_case)%expr @
+ (λ x2 : var (type.base base.type.unit),
+ UnderLets.to_expr (x ($x2)))%expr @
+ (λ (x2 : var (type.base A))(x3 : var
+ (type.base
+ (base.type.list
+ A))),
+ UnderLets.to_expr (x0 ($x2) ($x3)))%expr @ x1)%expr_pat)
end))%under_lets
- | None =>
- UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x2)))%expr @
- (λ (x2 : var (type.base A))(x3 : var
- (type.base
- (base.type.list A))),
- UnderLets.to_expr (x0 ($x2) ($x3)))%expr @ x1)%expr_pat
- end
- | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr _ |
- @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ =>
- UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x3 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x3)))%expr @
- (λ (x3 : var (type.base A))(x4 : var
- (type.base (base.type.list A))),
- UnderLets.to_expr (x0 ($x3) ($x4)))%expr @ x1)%expr_pat
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x3) x2 =>
- match match idc with
- | @ident.cons t0 => Some t0
- | _ => None
- end with
- | Some _ =>
- match
- s0 as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base P)))
- with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2) =>
- match
- s as t3
- return
- (Compile.value' false t3 ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base P)))
- with
- | type.base t3 =>
- fun v0 : defaults.expr (type.base t3) =>
- Compile.castbe v
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base P)))
- (fun x4 : option (defaults.expr (type.base A)) =>
- match x4 with
- | Some x5 =>
- Compile.castbe v0
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base P)))
- (fun
- xs : option
- (defaults.expr
- (type.base (base.type.list A))) =>
- match xs with
- | Some xs0 =>
+ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc) x3) x2 =>
+ _ <- pattern.ident.invert_bind_args idc pattern.ident.cons;
+ match
+ s0 as t2
+ return
+ (value' false t2 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base P))))
+ with
+ | type.base t2 =>
+ fun v : defaults.expr (type.base t2) =>
+ match
+ s as t3
+ return
+ (value' false t3 ->
+ option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base P))))
+ with
+ | type.base t3 =>
+ fun v0 : defaults.expr (type.base t3) =>
+ castbe v
+ (option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base P))))
+ (fun x4 : option (defaults.expr (type.base A)) =>
+ (x5 <- x4;
+ Some
+ (castbe v0
+ (option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base P))))
+ (fun
+ xs : option
+ (defaults.expr
+ (type.base (base.type.list A))) =>
+ (xs0 <- xs;
+ Some
+ (Some
(fv <-- (e <-- x0 x5 xs0;
UnderLets.Base
{|
anyexpr_ty := P;
unwrap := e |});
base.try_make_transport_cps
- (fun t0 : base.type =>
- defaults.expr (type.base t0))
(let (anyexpr_ty, _) := fv in anyexpr_ty)
P
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base P)))
(fun
a : option
(defaults.expr
@@ -1164,719 +819,607 @@ match idc in (ident t) return (Compile.value' true t) with
unwrap))
| None =>
UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x6 : var
- (type.base
- base.type.unit),
- UnderLets.to_expr (x ($x6)))%expr @
- (λ (x6 : var (type.base A))
- (x7 : var
- (type.base
- (base.type.list A))),
- UnderLets.to_expr
- (x0 ($x6) ($x7)))%expr @ x1)%expr_pat
- end))%under_lets
- | None =>
- UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x6 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x6)))%expr @
- (λ (x6 : var (type.base A))(x7 :
- var
- (type.base
- (base.type.list
- A))),
- UnderLets.to_expr (x0 ($x6) ($x7)))%expr @
- x1)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x5 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x5)))%expr @
- (λ (x5 : var (type.base A))(x6 : var
- (type.base
- (base.type.list
- A))),
- UnderLets.to_expr (x0 ($x5) ($x6)))%expr @ x1)%expr_pat
- end)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1
- =>
- UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x4 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x4)))%expr @
- (λ (x4 : var (type.base A))(x5 : var
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.list_case) @
+ ??{() -> ??} @
+ ??{?? -> ?? -> ??} @
+ (?? :: ??))
+ (#(ident.list_case)%expr @
+ (λ x6 : var
(type.base
- (base.type.list
- A))),
- UnderLets.to_expr (x0 ($x4) ($x5)))%expr @ x1)%expr_pat
- end (Compile.reflect x2)
- | (s1 -> d1)%ptype =>
- fun _ : Compile.value' false s1 -> Compile.value' true d1 =>
- UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x4 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x4)))%expr @
- (λ (x4 : var (type.base A))(x5 : var
- (type.base
- (base.type.list A))),
- UnderLets.to_expr (x0 ($x4) ($x5)))%expr @ x1)%expr_pat
- end (Compile.reflect x3)
- | None =>
- UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x4 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x4)))%expr @
- (λ (x4 : var (type.base A))(x5 : var
- (type.base
- (base.type.list A))),
- UnderLets.to_expr (x0 ($x4) ($x5)))%expr @ x1)%expr_pat
- end
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App _
- _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ =>
- UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x4 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x4)))%expr @
- (λ (x4 : var (type.base A))(x5 : var
- (type.base (base.type.list A))),
- UnderLets.to_expr (x0 ($x4) ($x5)))%expr @ x1)%expr_pat
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ |
- @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ =>
- UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x5 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x5)))%expr @
- (λ (x5 : var (type.base A))(x6 : var
- (type.base (base.type.list A))),
- UnderLets.to_expr (x0 ($x5) ($x6)))%expr @ x1)%expr_pat
- | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x4 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x4)))%expr @
- (λ (x4 : var (type.base A))(x5 : var
- (type.base (base.type.list A))),
- UnderLets.to_expr (x0 ($x4) ($x5)))%expr @ x1)%expr_pat
- | @expr.LetIn _ _ _ _ _ _ _ =>
- UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x3 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x3)))%expr @
- (λ (x3 : var (type.base A))(x4 : var
- (type.base (base.type.list A))),
- UnderLets.to_expr (x0 ($x3) ($x4)))%expr @ x1)%expr_pat
- | _ =>
- UnderLets.Base
- (#(ident.list_case)%expr @
- (λ x2 : var (type.base base.type.unit),
- UnderLets.to_expr (x ($x2)))%expr @
- (λ (x2 : var (type.base A))(x3 : var
- (type.base (base.type.list A))),
- UnderLets.to_expr (x0 ($x2) ($x3)))%expr @ x1)%expr_pat
- end
+ base.type.unit),
+ UnderLets.to_expr (x ($x6)))%expr @
+ (λ (x6 : var (type.base A))
+ (x7 : var
+ (type.base
+ (base.type.list A))),
+ UnderLets.to_expr
+ (x0 ($x6) ($x7)))%expr @
+ x1)%expr_pat)
+ end))%under_lets));;;
+ None)));;;
+ None)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x2)
+ | (s1 -> d1)%ptype =>
+ fun _ : value' false s1 -> value' true d1 => None
+ end (reflect x3)
+ | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
+ _ _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ |
+ @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ |
+ @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ => None
+ | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr
+ _ | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s
+ _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base
+ (#(ident.list_case)%expr @
+ (λ x2 : var (type.base base.type.unit),
+ UnderLets.to_expr (x ($x2)))%expr @
+ (λ (x2 : var (type.base A))(x3 : var (type.base (base.type.list A))),
+ UnderLets.to_expr (x0 ($x2) ($x3)))%expr @ x1)%expr_pat)%option
| @ident.List_length T =>
fun x : defaults.expr (type.base (base.type.list T)) =>
- reflect_list_cps x
- (fun xs : option (list (defaults.expr (type.base T))) =>
- match xs with
- | Some xs0 => UnderLets.Base ##(length xs0)%expr
- | None => UnderLets.Base (#(ident.List_length)%expr @ x)%expr_pat
- end)
+ (reflect_list_cps x
+ (fun xs : option (list (defaults.expr (type.base T))) =>
+ (xs0 <- xs;
+ Some (Some (UnderLets.Base (##(length xs0))%expr)));;;
+ None);;;
+ UnderLets.Base (#(ident.List_length)%expr @ x)%expr_pat)%option
| ident.List_seq =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralNat;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralNat;
+ Some
+ (UnderLets.Base
+ (fold_right
+ (fun (x1 : defaults.expr (type.base base.type.nat))
+ (xs : defaults.expr
+ (type.base (base.type.list base.type.nat)))
+ => (x1 :: xs)%expr_pat) []%expr_pat
+ (map (fun v : nat => (##v)%expr) (seq args args0))))
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option nat)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args0 =>
- UnderLets.Base
- (fold_right
- (fun (x1 : defaults.expr (type.base base.type.nat))
- (xs : defaults.expr
- (type.base (base.type.list base.type.nat)))
- => (x1 :: xs)%expr_pat) []%expr_pat
- (map (fun v : nat => ##(v)%expr) (seq args args0)))
- | None =>
- UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat
- end
- | None => UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat)%option
| @ident.List_firstn A =>
fun (x : defaults.expr (type.base base.type.nat))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args =>
- reflect_list_cps x0
- (fun xs : option (list (defaults.expr (type.base A))) =>
- match xs with
- | Some xs0 =>
- base.try_make_transport_cps
- (fun A0 : base.type =>
- defaults.expr (type.base (base.type.list A0))) A A
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base (base.type.list A))))
- (fun
- a : option
- (defaults.expr (type.base (base.type.list A)) ->
- defaults.expr (type.base (base.type.list A)))
- =>
- match a with
- | Some x' =>
- UnderLets.Base (x' (reify_list (firstn args xs0)))
- | None =>
- UnderLets.Base
- (#(ident.List_firstn)%expr @ x @ x0)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.List_firstn)%expr @ x @ x0)%expr_pat
- end)
- | None =>
- UnderLets.Base (#(ident.List_firstn)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.List_firstn)%expr @ x @ x0)%expr_pat
- end
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralNat;
+ reflect_list_cps x0
+ (fun xs : option (list (defaults.expr (type.base A))) =>
+ (xs0 <- xs;
+ Some
+ (Some
+ (base.try_make_transport_cps A A
+ (fun
+ a : option
+ (defaults.expr (type.base (base.type.list A)) ->
+ defaults.expr (type.base (base.type.list A)))
+ =>
+ match a with
+ | Some x' =>
+ UnderLets.Base (x' (reify_list (firstn args xs0)))
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.List_firstn) @
+ #(pattern.ident.LiteralNat) @
+ ??{type.base (pattern.base.type.list ??)})
+ (#(ident.List_firstn)%expr @ x @ x0)%expr_pat)
+ end))));;;
+ None)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.List_firstn)%expr @ x @ x0)%expr_pat)%option
| @ident.List_skipn A =>
fun (x : defaults.expr (type.base base.type.nat))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args =>
- reflect_list_cps x0
- (fun xs : option (list (defaults.expr (type.base A))) =>
- match xs with
- | Some xs0 =>
- base.try_make_transport_cps
- (fun A0 : base.type =>
- defaults.expr (type.base (base.type.list A0))) A A
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base (base.type.list A))))
- (fun
- a : option
- (defaults.expr (type.base (base.type.list A)) ->
- defaults.expr (type.base (base.type.list A)))
- =>
- match a with
- | Some x' =>
- UnderLets.Base (x' (reify_list (skipn args xs0)))
- | None =>
- UnderLets.Base
- (#(ident.List_skipn)%expr @ x @ x0)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.List_skipn)%expr @ x @ x0)%expr_pat
- end)
- | None => UnderLets.Base (#(ident.List_skipn)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.List_skipn)%expr @ x @ x0)%expr_pat
- end
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralNat;
+ reflect_list_cps x0
+ (fun xs : option (list (defaults.expr (type.base A))) =>
+ (xs0 <- xs;
+ Some
+ (Some
+ (base.try_make_transport_cps A A
+ (fun
+ a : option
+ (defaults.expr (type.base (base.type.list A)) ->
+ defaults.expr (type.base (base.type.list A)))
+ =>
+ match a with
+ | Some x' =>
+ UnderLets.Base (x' (reify_list (skipn args xs0)))
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.List_skipn) @
+ #(pattern.ident.LiteralNat) @
+ ??{type.base (pattern.base.type.list ??)})
+ (#(ident.List_skipn)%expr @ x @ x0)%expr_pat)
+ end))));;;
+ None)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.List_skipn)%expr @ x @ x0)%expr_pat)%option
| @ident.List_repeat A =>
fun (x : defaults.expr (type.base A))
(x0 : defaults.expr (type.base base.type.nat)) =>
- match x0 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args =>
- base.try_make_transport_cps
- (fun A0 : base.type =>
- defaults.expr (type.base (base.type.list A0))) A A
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base (base.type.list A))))
- (fun
- a : option
- (defaults.expr (type.base (base.type.list A)) ->
- defaults.expr (type.base (base.type.list A))) =>
- match a with
- | Some x' => UnderLets.Base (x' (reify_list (repeat x args)))
- | None =>
- UnderLets.Base
- (#(ident.List_repeat)%expr @ x @ x0)%expr_pat
- end)
- | None =>
- UnderLets.Base (#(ident.List_repeat)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.List_repeat)%expr @ x @ x0)%expr_pat
- end
+ ((match x0 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralNat;
+ Some
+ (base.try_make_transport_cps A A
+ (fun
+ a : option
+ (defaults.expr (type.base (base.type.list A)) ->
+ defaults.expr (type.base (base.type.list A))) =>
+ match a with
+ | Some x' => UnderLets.Base (x' (reify_list (repeat x args)))
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.List_repeat) @ ?? @
+ #(pattern.ident.LiteralNat))
+ (#(ident.List_repeat)%expr @ x @ x0)%expr_pat)
+ end))
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.List_repeat)%expr @ x @ x0)%expr_pat)%option
| @ident.List_combine A B =>
fun (x : defaults.expr (type.base (base.type.list A)))
(x0 : defaults.expr (type.base (base.type.list B))) =>
- reflect_list_cps x
- (fun xs : option (list (defaults.expr (type.base A))) =>
- match xs with
- | Some xs0 =>
- reflect_list_cps x0
- (fun ys : option (list (defaults.expr (type.base B))) =>
- match ys with
- | Some ys0 =>
- (trA <-- base.try_make_transport_cps
- (fun A0 : base.type =>
- defaults.expr
- (type.base (base.type.list (A0 * B)))) A A;
- trB <-- base.try_make_transport_cps
- (fun B0 : base.type =>
- defaults.expr
- (type.base (base.type.list (A * B0)))) B B;
- return Some
- (fun
- v : defaults.expr
- (type.base (base.type.list (A * B))) =>
- trB (trA v)))%cps
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base (base.type.list (A * B)))))
- (fun
- a : option
- (defaults.expr
- (type.base (base.type.list (A * B))) ->
- defaults.expr
- (type.base (base.type.list (A * B)))) =>
- match a with
- | Some x' =>
- UnderLets.Base
- (x'
- (reify_list
- (map (fun '(x1, y) => (x1, y)%expr_pat)
- (combine xs0 ys0))))
- | None =>
- UnderLets.Base
- (#(ident.List_combine)%expr @ x @ x0)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.List_combine)%expr @ x @ x0)%expr_pat
- end)
- | None =>
- UnderLets.Base (#(ident.List_combine)%expr @ x @ x0)%expr_pat
- end)
+ (reflect_list_cps x
+ (fun xs : option (list (defaults.expr (type.base A))) =>
+ (xs0 <- xs;
+ Some
+ (reflect_list_cps x0
+ (fun ys : option (list (defaults.expr (type.base B))) =>
+ (ys0 <- ys;
+ Some
+ (Some
+ ((trA <-- @base.try_make_transport_cps
+ (fun A0 : base.type =>
+ defaults.expr
+ (type.base (base.type.list (A0 * B)))) A
+ A;
+ trB <-- @base.try_make_transport_cps
+ (fun B0 : base.type =>
+ defaults.expr
+ (type.base (base.type.list (A * B0)))) B
+ B;
+ return Some
+ (fun
+ v : defaults.expr
+ (type.base (base.type.list (A * B)))
+ => trB (trA v)))%cps
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.list (A * B)))))
+ (fun
+ a : option
+ (defaults.expr
+ (type.base (base.type.list (A * B))) ->
+ defaults.expr
+ (type.base (base.type.list (A * B)))) =>
+ match a with
+ | Some x' =>
+ UnderLets.Base
+ (x'
+ (reify_list
+ (map (fun '(x1, y) => (x1, y)%expr_pat)
+ (combine xs0 ys0))))
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.List_combine) @
+ ??{type.base (pattern.base.type.list ??)} @
+ ??{type.base (pattern.base.type.list ??)})
+ (#(ident.List_combine)%expr @ x @ x0)%expr_pat)
+ end))));;;
+ None)));;;
+ None);;;
+ UnderLets.Base (#(ident.List_combine)%expr @ x @ x0)%expr_pat)%option
| @ident.List_map A B =>
fun
(x : defaults.expr (type.base A) ->
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base B)))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- Compile.castbe x0
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base (base.type.list B))))
- (fun e : option (defaults.expr (type.base (base.type.list A))) =>
- match e with
- | Some e0 =>
- reflect_list_cps e0
- (fun ls : option (list (defaults.expr (type.base A))) =>
- match ls with
- | Some ls0 =>
- (fv <-- (e1 <-- list_rect
- (fun
- _ : list (defaults.expr (type.base A))
- =>
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.list B))))
- (UnderLets.Base []%expr_pat)
- (fun (x1 : defaults.expr (type.base A))
- (_ : list
- (defaults.expr (type.base A)))
- (rec : UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.list B))))
- =>
- rec' <-- rec;
- fx <-- x x1;
- UnderLets.Base (fx :: rec')%expr_pat)
- ls0;
- UnderLets.Base
- {|
- anyexpr_ty := base.type.list B;
- unwrap := e1 |});
- base.try_make_transport_cps
- (fun t : base.type => defaults.expr (type.base t))
- (let (anyexpr_ty, _) := fv in anyexpr_ty)
- (base.type.list B)
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base (base.type.list B))))
- (fun
- a : option
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := fv in anyexpr_ty)) ->
- defaults.expr (type.base (base.type.list B)))
- =>
- match a with
- | Some x' =>
- UnderLets.Base
- (x'
- (let
- (anyexpr_ty, unwrap) as a0
- return
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := a0 in
- anyexpr_ty))) := fv in
- unwrap))
- | None =>
- UnderLets.Base
- (#(ident.List_map)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
- end))%under_lets
- | None =>
- UnderLets.Base
- (#(ident.List_map)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.List_map)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
- end)
+ ((castbe x0
+ (option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base (base.type.list B)))))
+ (fun e : option (defaults.expr (type.base (base.type.list A))) =>
+ (e0 <- e;
+ Some
+ (reflect_list_cps e0
+ (fun ls : option (list (defaults.expr (type.base A))) =>
+ (ls0 <- ls;
+ Some
+ (Some
+ (fv <-- (e1 <-- list_rect
+ (fun
+ _ : list
+ (defaults.expr (type.base A))
+ =>
+ UnderLets.UnderLets base.type ident
+ var
+ (defaults.expr
+ (type.base (base.type.list B))))
+ (UnderLets.Base []%expr_pat)
+ (fun
+ (x1 : defaults.expr (type.base A))
+ (_ : list
+ (defaults.expr
+ (type.base A)))
+ (rec : UnderLets.UnderLets
+ base.type ident var
+ (defaults.expr
+ (type.base
+ (base.type.list B))))
+ =>
+ rec' <-- rec;
+ fx <-- x x1;
+ UnderLets.Base (fx :: rec')%expr_pat)
+ ls0;
+ UnderLets.Base
+ {|
+ anyexpr_ty := base.type.list B;
+ unwrap := e1 |});
+ base.try_make_transport_cps
+ (let (anyexpr_ty, _) := fv in anyexpr_ty)
+ (base.type.list B)
+ (fun
+ a : option
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) := fv in
+ anyexpr_ty)) ->
+ defaults.expr
+ (type.base (base.type.list B))) =>
+ match a with
+ | Some x' =>
+ UnderLets.Base
+ (x'
+ (let
+ (anyexpr_ty, unwrap) as a0
+ return
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) := a0 in
+ anyexpr_ty))) := fv in
+ unwrap))
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.List_map) @ ??{?? -> ??} @
+ ??{type.base (pattern.base.type.list ??)})
+ (#(ident.List_map)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat)
+ end))%under_lets));;;
+ None)));;;
+ None);;
+ None);;;
+ UnderLets.Base
+ (#(ident.List_map)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat)%option
| @ident.List_app A =>
fun x x0 : defaults.expr (type.base (base.type.list A)) =>
- Compile.castbe x
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base (base.type.list A))))
- (fun e : option (defaults.expr (type.base (base.type.list A))) =>
- match e with
- | Some e0 =>
- reflect_list_cps e0
- (fun ls : option (list (defaults.expr (type.base A))) =>
- match ls with
- | Some ls0 =>
- (fv <-- (e1 <-- list_rect
- (fun
- _ : list (defaults.expr (type.base A))
- =>
- UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.list A))))
- (UnderLets.Base x0)
- (fun (x1 : defaults.expr (type.base A))
- (_ : list
- (defaults.expr (type.base A)))
- (rec : UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.list A))))
- =>
- rec' <-- rec;
- UnderLets.Base (x1 :: rec')%expr_pat)
- ls0;
- UnderLets.Base
- {|
- anyexpr_ty := base.type.list A;
- unwrap := e1 |});
- base.try_make_transport_cps
- (fun t : base.type => defaults.expr (type.base t))
- (let (anyexpr_ty, _) := fv in anyexpr_ty)
- (base.type.list A)
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base (base.type.list A))))
- (fun
- a : option
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := fv in anyexpr_ty)) ->
- defaults.expr (type.base (base.type.list A)))
- =>
- match a with
- | Some x' =>
- UnderLets.Base
- (x'
- (let
- (anyexpr_ty, unwrap) as a0
- return
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := a0 in
- anyexpr_ty))) := fv in
- unwrap))
- | None => UnderLets.Base (x ++ x0)%expr
- end))%under_lets
- | None => UnderLets.Base (x ++ x0)%expr
- end)
- | None => UnderLets.Base (x ++ x0)%expr
- end)
+ ((castbe x
+ (option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base (base.type.list A)))))
+ (fun e : option (defaults.expr (type.base (base.type.list A))) =>
+ (e0 <- e;
+ Some
+ (reflect_list_cps e0
+ (fun ls : option (list (defaults.expr (type.base A))) =>
+ (ls0 <- ls;
+ Some
+ (Some
+ (fv <-- (e1 <-- list_rect
+ (fun
+ _ : list
+ (defaults.expr (type.base A))
+ =>
+ UnderLets.UnderLets base.type ident
+ var
+ (defaults.expr
+ (type.base (base.type.list A))))
+ (UnderLets.Base x0)
+ (fun
+ (x1 : defaults.expr (type.base A))
+ (_ : list
+ (defaults.expr
+ (type.base A)))
+ (rec : UnderLets.UnderLets
+ base.type ident var
+ (defaults.expr
+ (type.base
+ (base.type.list A))))
+ =>
+ rec' <-- rec;
+ UnderLets.Base (x1 :: rec')%expr_pat)
+ ls0;
+ UnderLets.Base
+ {|
+ anyexpr_ty := base.type.list A;
+ unwrap := e1 |});
+ base.try_make_transport_cps
+ (let (anyexpr_ty, _) := fv in anyexpr_ty)
+ (base.type.list A)
+ (fun
+ a : option
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) := fv in
+ anyexpr_ty)) ->
+ defaults.expr
+ (type.base (base.type.list A))) =>
+ match a with
+ | Some x' =>
+ UnderLets.Base
+ (x'
+ (let
+ (anyexpr_ty, unwrap) as a0
+ return
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) := a0 in
+ anyexpr_ty))) := fv in
+ unwrap))
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (??{type.base (pattern.base.type.list ??)} ++
+ ??{type.base (pattern.base.type.list ??)})
+ (x ++ x0))
+ end))%under_lets));;;
+ None)));;;
+ None);;
+ None);;;
+ UnderLets.Base (x ++ x0)%expr)%option
| @ident.List_rev A =>
fun x : defaults.expr (type.base (base.type.list A)) =>
- reflect_list_cps x
- (fun xs : option (list (defaults.expr (type.base A))) =>
- match xs with
- | Some xs0 =>
- base.try_make_transport_cps
- (fun A0 : base.type =>
- defaults.expr (type.base (base.type.list A0))) A A
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base (base.type.list A))))
- (fun
- a : option
- (defaults.expr (type.base (base.type.list A)) ->
- defaults.expr (type.base (base.type.list A))) =>
- match a with
- | Some x' => UnderLets.Base (x' (reify_list (rev xs0)))
- | None => UnderLets.Base (#(ident.List_rev)%expr @ x)%expr_pat
- end)
- | None => UnderLets.Base (#(ident.List_rev)%expr @ x)%expr_pat
- end)
+ ((reflect_list_cps x
+ (fun xs : option (list (defaults.expr (type.base A))) =>
+ (xs0 <- xs;
+ Some
+ (Some
+ (base.try_make_transport_cps A A
+ (fun
+ a : option
+ (defaults.expr (type.base (base.type.list A)) ->
+ defaults.expr (type.base (base.type.list A))) =>
+ match a with
+ | Some x' => UnderLets.Base (x' (reify_list (rev xs0)))
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.List_rev) @
+ ??{type.base (pattern.base.type.list ??)})
+ (#(ident.List_rev)%expr @ x)%expr_pat)
+ end))));;;
+ None);;
+ None);;;
+ UnderLets.Base (#(ident.List_rev)%expr @ x)%expr_pat)%option
| @ident.List_flat_map A B =>
fun
(x : defaults.expr (type.base A) ->
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base (base.type.list B))))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- Compile.castbe x0
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base (base.type.list B))))
- (fun e : option (defaults.expr (type.base (base.type.list A))) =>
- match e with
- | Some e0 =>
- reflect_list_cps e0
- (fun ls : option (list (defaults.expr (type.base A))) =>
- match ls with
- | Some ls0 =>
- (fv <-- (e1 <-- list_rect
- (fun
- _ : list (defaults.expr (type.base A))
- =>
- UnderLets.UnderLets base.type ident var
+ ((castbe x0
+ (option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base (base.type.list B)))))
+ (fun e : option (defaults.expr (type.base (base.type.list A))) =>
+ (e0 <- e;
+ Some
+ (reflect_list_cps e0
+ (fun ls : option (list (defaults.expr (type.base A))) =>
+ (ls0 <- ls;
+ Some
+ (Some
+ (fv <-- (e1 <-- list_rect
+ (fun
+ _ : list
+ (defaults.expr (type.base A))
+ =>
+ UnderLets.UnderLets base.type ident
+ var
+ (defaults.expr
+ (type.base (base.type.list B))))
+ (UnderLets.Base []%expr_pat)
+ (fun
+ (x1 : defaults.expr (type.base A))
+ (_ : list
+ (defaults.expr
+ (type.base A)))
+ (rec : UnderLets.UnderLets
+ base.type ident var
+ (defaults.expr
+ (type.base
+ (base.type.list B))))
+ =>
+ rec' <-- rec;
+ fx <-- x x1;
+ UnderLets.Base ($fx ++ rec')%expr)
+ ls0;
+ UnderLets.Base
+ {|
+ anyexpr_ty := base.type.list B;
+ unwrap := e1 |});
+ fv0 <-- do_again
+ (let (anyexpr_ty, _) := fv in anyexpr_ty)
+ (let
+ (anyexpr_ty, unwrap) as a
+ return
(defaults.expr
- (type.base (base.type.list B))))
- (UnderLets.Base []%expr_pat)
- (fun (x1 : defaults.expr (type.base A))
- (_ : list
- (defaults.expr (type.base A)))
- (rec : UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.list B))))
- =>
- rec' <-- rec;
- fx <-- x x1;
- UnderLets.Base ($fx ++ rec')%expr) ls0;
- UnderLets.Base
- {|
- anyexpr_ty := base.type.list B;
- unwrap := e1 |});
- fv0 <-- do_again (let (anyexpr_ty, _) := fv in anyexpr_ty)
- (let
- (anyexpr_ty, unwrap) as a
- return
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := a in
- anyexpr_ty))) := fv in
- unwrap);
- base.try_make_transport_cps
- (fun t : base.type => defaults.expr (type.base t))
- (let (anyexpr_ty, _) := fv in anyexpr_ty)
- (base.type.list B)
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base (base.type.list B))))
- (fun
- a : option
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := fv in anyexpr_ty)) ->
- defaults.expr (type.base (base.type.list B)))
- =>
- match a with
- | Some x' => UnderLets.Base (x' fv0)
- | None =>
- UnderLets.Base
- (#(ident.List_flat_map)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
- end))%under_lets
- | None =>
- UnderLets.Base
- (#(ident.List_flat_map)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.List_flat_map)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
- end)
+ (type.base
+ (let (anyexpr_ty, _) := a in
+ anyexpr_ty))) := fv in
+ unwrap);
+ base.try_make_transport_cps
+ (let (anyexpr_ty, _) := fv in anyexpr_ty)
+ (base.type.list B)
+ (fun
+ a : option
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) := fv in
+ anyexpr_ty)) ->
+ defaults.expr
+ (type.base (base.type.list B))) =>
+ match a with
+ | Some x' => UnderLets.Base (x' fv0)
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.List_flat_map) @
+ ??{?? ->
+ type.base (pattern.base.type.list ??)} @
+ ??{type.base (pattern.base.type.list ??)})
+ (#(ident.List_flat_map)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat)
+ end))%under_lets));;;
+ None)));;;
+ None);;
+ None);;;
+ UnderLets.Base
+ (#(ident.List_flat_map)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat)%option
| @ident.List_partition A =>
fun
(x : defaults.expr (type.base A) ->
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base base.type.bool)))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- Compile.castbe x0
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base (base.type.list A * base.type.list A)%etype)))
- (fun e : option (defaults.expr (type.base (base.type.list A))) =>
- match e with
- | Some e0 =>
- reflect_list_cps e0
- (fun ls : option (list (defaults.expr (type.base A))) =>
- match ls with
- | Some ls0 =>
- (fv <-- (e1 <-- list_rect
- (fun
- _ : list (defaults.expr (type.base A))
- =>
- UnderLets.UnderLets base.type ident var
+ ((castbe x0
+ (option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr
+ (type.base (base.type.list A * base.type.list A)%etype))))
+ (fun e : option (defaults.expr (type.base (base.type.list A))) =>
+ (e0 <- e;
+ Some
+ (reflect_list_cps e0
+ (fun ls : option (list (defaults.expr (type.base A))) =>
+ (ls0 <- ls;
+ Some
+ (Some
+ (fv <-- (e1 <-- list_rect
+ (fun
+ _ : list
+ (defaults.expr (type.base A))
+ =>
+ UnderLets.UnderLets base.type ident
+ var
+ (defaults.expr
+ (type.base
+ (base.type.list A *
+ base.type.list A)%etype)))
+ (UnderLets.Base ([], [])%expr_pat)
+ (fun
+ (x1 : defaults.expr (type.base A))
+ (_ : list
+ (defaults.expr
+ (type.base A)))
+ (rec : UnderLets.UnderLets
+ base.type ident var
+ (defaults.expr
+ (type.base
+ (base.type.list A *
+ base.type.list A)%etype)))
+ =>
+ rec' <-- rec;
+ fx <-- id x x1;
+ UnderLets.Base
+ (#(ident.prod_rect)%expr @
+ (λ g
+ d : defaults.expr
+ (type.base
+ (base.type.list A)),
+ (#(ident.bool_rect)%expr @
+ (λ _ : defaults.expr
+ (type.base
+ base.type.unit),
+ ($x1 :: $g, $d)%expr_pat) @
+ (λ _ : defaults.expr
+ (type.base
+ base.type.unit),
+ ($g, $x1 :: $d)%expr_pat) @
+ $fx)%expr_pat)%expr @ rec')%expr_pat)
+ ls0;
+ UnderLets.Base
+ {|
+ anyexpr_ty := (base.type.list A *
+ base.type.list A)%etype;
+ unwrap := e1 |});
+ fv0 <-- do_again
+ (let (anyexpr_ty, _) := fv in anyexpr_ty)
+ (let
+ (anyexpr_ty, unwrap) as a
+ return
(defaults.expr
(type.base
- (base.type.list A *
- base.type.list A)%etype)))
- (UnderLets.Base ([], [])%expr_pat)
- (fun (x1 : defaults.expr (type.base A))
- (_ : list
- (defaults.expr (type.base A)))
- (rec : UnderLets.UnderLets base.type
- ident var
- (defaults.expr
- (type.base
- (base.type.list A *
- base.type.list A)%etype)))
- =>
- rec' <-- rec;
- fx <-- id x x1;
- UnderLets.Base
- (#(ident.prod_rect)%expr @
- (λ g
- d : defaults.expr
- (type.base
- (base.type.list A)),
- (#(ident.bool_rect)%expr @
- (λ _ : defaults.expr
- (type.base base.type.unit),
- ($x1 :: $g, $d)%expr_pat) @
- (λ _ : defaults.expr
- (type.base base.type.unit),
- ($g, $x1 :: $d)%expr_pat) @ $fx)%expr_pat)%expr @
- rec')%expr_pat) ls0;
- UnderLets.Base
- {|
- anyexpr_ty := (base.type.list A *
- base.type.list A)%etype;
- unwrap := e1 |});
- fv0 <-- do_again (let (anyexpr_ty, _) := fv in anyexpr_ty)
- (let
- (anyexpr_ty, unwrap) as a
- return
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := a in
- anyexpr_ty))) := fv in
- unwrap);
- base.try_make_transport_cps
- (fun t : base.type => defaults.expr (type.base t))
- (let (anyexpr_ty, _) := fv in anyexpr_ty)
- (base.type.list A * base.type.list A)
- (UnderLets.UnderLets base.type ident var
- (defaults.expr
- (type.base
- (base.type.list A * base.type.list A)%etype)))
- (fun
- a : option
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := fv in anyexpr_ty)) ->
- defaults.expr
- (type.base
- (base.type.list A * base.type.list A)%etype))
- =>
- match a with
- | Some x' => UnderLets.Base (x' fv0)
- | None =>
- UnderLets.Base
- (#(ident.List_partition)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
- end))%under_lets
- | None =>
- UnderLets.Base
- (#(ident.List_partition)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.List_partition)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
- end)
+ (let (anyexpr_ty, _) := a in
+ anyexpr_ty))) := fv in
+ unwrap);
+ base.try_make_transport_cps
+ (let (anyexpr_ty, _) := fv in anyexpr_ty)
+ (base.type.list A * base.type.list A)
+ (fun
+ a : option
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) := fv in
+ anyexpr_ty)) ->
+ defaults.expr
+ (type.base
+ (base.type.list A * base.type.list A)%etype))
+ =>
+ match a with
+ | Some x' => UnderLets.Base (x' fv0)
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.List_partition) @
+ ??{?? -> type.base base.type.bool} @
+ ??{type.base (pattern.base.type.list ??)})
+ (#(ident.List_partition)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat)
+ end))%under_lets));;;
+ None)));;;
+ None);;
+ None);;;
+ UnderLets.Base
+ (#(ident.List_partition)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat)%option
| @ident.List_fold_right A B =>
fun
(x : defaults.expr (type.base B) ->
@@ -1885,3239 +1428,1230 @@ match idc in (ident t) return (Compile.value' true t) with
(defaults.expr (type.base A)))
(x0 : defaults.expr (type.base A))
(x1 : defaults.expr (type.base (base.type.list B))) =>
- Compile.castv x
- (UnderLets.UnderLets base.type ident var (defaults.expr (type.base A)))
- (fun
- f : option
- (defaults.expr (type.base B) ->
- defaults.expr (type.base A) ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base A))) =>
- match f with
- | Some f0 =>
- reflect_list_cps x1
- (fun ls : option (list (defaults.expr (type.base B))) =>
- match ls with
- | Some ls0 =>
- (fv <-- (e <-- list_rect
- (fun
- _ : list (defaults.expr (type.base B))
- =>
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base A)))
- (UnderLets.Base x0)
- (fun (x2 : defaults.expr (type.base B))
- (_ : list (defaults.expr (type.base B)))
- (rec : UnderLets.UnderLets base.type
- ident var
- (defaults.expr (type.base A)))
- => rec' <-- rec;
- f0 x2 rec') ls0;
- UnderLets.Base {| anyexpr_ty := A; unwrap := e |});
- base.try_make_transport_cps
- (fun t : base.type => defaults.expr (type.base t))
- (let (anyexpr_ty, _) := fv in anyexpr_ty) A
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base A)))
- (fun
- a : option
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := fv in anyexpr_ty)) ->
- defaults.expr (type.base A)) =>
- match a with
- | Some x' =>
- UnderLets.Base
- (x'
- (let
- (anyexpr_ty, unwrap) as a0
- return
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := a0 in
- anyexpr_ty))) := fv in
- unwrap))
- | None =>
- UnderLets.Base
- (#(ident.List_fold_right)%expr @
- (λ (x2 : var (type.base B))(x3 : var
- (type.base A)),
- UnderLets.to_expr (x ($x2) ($x3)))%expr @ x0 @
- x1)%expr_pat
- end))%under_lets
- | None =>
- UnderLets.Base
- (#(ident.List_fold_right)%expr @
- (λ (x2 : var (type.base B))(x3 : var (type.base A)),
- UnderLets.to_expr (x ($x2) ($x3)))%expr @ x0 @ x1)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.List_fold_right)%expr @
- (λ (x2 : var (type.base B))(x3 : var (type.base A)),
- UnderLets.to_expr (x ($x2) ($x3)))%expr @ x0 @ x1)%expr_pat
- end)
+ ((castv x
+ (option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base A))))
+ (fun
+ f : option
+ (defaults.expr (type.base B) ->
+ defaults.expr (type.base A) ->
+ UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base A))) =>
+ (f0 <- f;
+ Some
+ (reflect_list_cps x1
+ (fun ls : option (list (defaults.expr (type.base B))) =>
+ (ls0 <- ls;
+ Some
+ (Some
+ (fv <-- (e <-- list_rect
+ (fun
+ _ : list
+ (defaults.expr (type.base B))
+ =>
+ UnderLets.UnderLets base.type ident
+ var (defaults.expr (type.base A)))
+ (UnderLets.Base x0)
+ (fun
+ (x2 : defaults.expr (type.base B))
+ (_ : list
+ (defaults.expr (type.base B)))
+ (rec : UnderLets.UnderLets
+ base.type ident var
+ (defaults.expr
+ (type.base A))) =>
+ rec' <-- rec;
+ f0 x2 rec') ls0;
+ UnderLets.Base
+ {| anyexpr_ty := A; unwrap := e |});
+ base.try_make_transport_cps
+ (let (anyexpr_ty, _) := fv in anyexpr_ty) A
+ (fun
+ a : option
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) := fv in
+ anyexpr_ty)) ->
+ defaults.expr (type.base A)) =>
+ match a with
+ | Some x' =>
+ UnderLets.Base
+ (x'
+ (let
+ (anyexpr_ty, unwrap) as a0
+ return
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) := a0 in
+ anyexpr_ty))) := fv in
+ unwrap))
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.List_fold_right) @
+ ??{?? -> ?? -> ??} @ ?? @
+ ??{type.base (pattern.base.type.list ??)})
+ (#(ident.List_fold_right)%expr @
+ (λ (x2 : var (type.base B))(x3 :
+ var
+ (type.base
+ A)),
+ UnderLets.to_expr (x ($x2) ($x3)))%expr @
+ x0 @ x1)%expr_pat)
+ end))%under_lets));;;
+ None)));;;
+ None);;
+ None);;;
+ UnderLets.Base
+ (#(ident.List_fold_right)%expr @
+ (λ (x2 : var (type.base B))(x3 : var (type.base A)),
+ UnderLets.to_expr (x ($x2) ($x3)))%expr @ x0 @ x1)%expr_pat)%option
| @ident.List_update_nth T =>
fun (x : defaults.expr (type.base base.type.nat))
(x0 : defaults.expr (type.base T) ->
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base T)))
(x1 : defaults.expr (type.base (base.type.list T))) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args =>
- Compile.castv x0
+ (match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralNat;
+ castv x0
+ (option
(UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base (base.type.list T))))
- (fun
- f : option
- (defaults.expr (type.base T) ->
- UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base T))) =>
- match f with
- | Some f0 =>
- reflect_list_cps x1
- (fun ls : option (list (defaults.expr (type.base T))) =>
- match ls with
- | Some ls0 =>
- (fv <-- (e <-- (retv <---- update_nth args
- (fun
- x2 : UnderLets.UnderLets
- base.type
- ident var
- (defaults.expr
- (type.base
+ (defaults.expr (type.base (base.type.list T)))))
+ (fun
+ f : option
+ (defaults.expr (type.base T) ->
+ UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base T))) =>
+ (f0 <- f;
+ Some
+ (reflect_list_cps x1
+ (fun ls : option (list (defaults.expr (type.base T))) =>
+ (ls0 <- ls;
+ Some
+ (Some
+ (fv <-- (e <-- (retv <---- update_nth args
+ (fun
+ x2 : UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
T)) =>
- x3 <-- x2;
- f0 x3)
- (map UnderLets.Base
- ls0);
- UnderLets.Base (reify_list retv));
- UnderLets.Base
- {|
- anyexpr_ty := base.type.list T;
- unwrap := e |});
- base.try_make_transport_cps
- (fun t0 : base.type =>
- defaults.expr (type.base t0))
- (let (anyexpr_ty, _) := fv in anyexpr_ty)
- (base.type.list T)
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base (base.type.list T))))
- (fun
- a : option
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) := fv in
- anyexpr_ty)) ->
- defaults.expr
- (type.base (base.type.list T))) =>
- match a with
- | Some x' =>
+ x3 <-- x2;
+ f0 x3)
+ (map UnderLets.Base ls0);
+ UnderLets.Base (reify_list retv));
UnderLets.Base
- (x'
- (let
- (anyexpr_ty, unwrap) as a0
- return
- (defaults.expr
- (type.base
- (let (anyexpr_ty, _) :=
- a0 in
- anyexpr_ty))) := fv in
- unwrap))
- | None =>
- UnderLets.Base
- (#(ident.List_update_nth)%expr @ x @
- (λ x2 : var (type.base T),
- UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat
- end))%under_lets
- | None =>
- UnderLets.Base
- (#(ident.List_update_nth)%expr @ x @
- (λ x2 : var (type.base T),
- UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.List_update_nth)%expr @ x @
- (λ x2 : var (type.base T),
- UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.List_update_nth)%expr @ x @
- (λ x2 : var (type.base T),
- UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat
- end
- | ($_)%expr | @expr.Abs _ _ _ _ _ _ =>
- UnderLets.Base
- (#(ident.List_update_nth)%expr @ x @
- (λ x2 : var (type.base T),
- UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat
- | _ =>
- UnderLets.Base
- (#(ident.List_update_nth)%expr @ x @
- (λ x3 : var (type.base T),
- UnderLets.to_expr (x0 ($x3)))%expr @ x1)%expr_pat
- end
-| @ident.List_nth_default T =>
- fun (x : defaults.expr (type.base T))
- (x0 : defaults.expr (type.base (base.type.list T)))
- (x1 : defaults.expr (type.base base.type.nat)) =>
- match x1 with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args =>
- Compile.castbe x
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base T)))
- (fun default : option (defaults.expr (type.base T)) =>
- match default with
- | Some default0 =>
- reflect_list_cps x0
- (fun ls : option (list (defaults.expr (type.base T))) =>
- match ls with
- | Some ls0 =>
+ {|
+ anyexpr_ty := base.type.list T;
+ unwrap := e |});
base.try_make_transport_cps
- (fun t0 : base.type =>
- defaults.expr (type.base t0)) T T
- (UnderLets.UnderLets base.type ident var
- (defaults.expr (type.base T)))
+ (let (anyexpr_ty, _) := fv in anyexpr_ty)
+ (base.type.list T)
(fun
a : option
- (defaults.expr (type.base T) ->
- defaults.expr (type.base T)) =>
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) := fv in
+ anyexpr_ty)) ->
+ defaults.expr
+ (type.base (base.type.list T))) =>
match a with
| Some x' =>
UnderLets.Base
- (x' (nth_default default0 ls0 args))
+ (x'
+ (let
+ (anyexpr_ty, unwrap) as a0
+ return
+ (defaults.expr
+ (type.base
+ (let (anyexpr_ty, _) :=
+ a0 in
+ anyexpr_ty))) := fv in
+ unwrap))
| None =>
UnderLets.Base
- (#(ident.List_nth_default)%expr @ x @ x0 @
- x1)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.List_nth_default)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.List_nth_default)%expr @ x @ x0 @ x1)%expr_pat
- end)
- | None =>
- UnderLets.Base
- (#(ident.List_nth_default)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.List_nth_default)%expr @ x @ x0 @ x1)%expr_pat
- end
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.List_update_nth) @
+ #(pattern.ident.LiteralNat) @
+ ??{?? -> ??} @
+ ??{type.base
+ (pattern.base.type.list ??)})
+ (#(ident.List_update_nth)%expr @ x @
+ (λ x2 : var (type.base T),
+ UnderLets.to_expr (x0 ($x2)))%expr @
+ x1)%expr_pat)
+ end))%under_lets));;;
+ None)));;;
+ None)
+ | _ => None
+ end;;;
+ UnderLets.Base
+ (#(ident.List_update_nth)%expr @ x @
+ (λ x2 : var (type.base T),
+ UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat)%option
+| @ident.List_nth_default T =>
+ fun (x : defaults.expr (type.base T))
+ (x0 : defaults.expr (type.base (base.type.list T)))
+ (x1 : defaults.expr (type.base base.type.nat)) =>
+ ((match x1 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pattern.ident.LiteralNat;
+ castbe x
+ (option
+ (UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base T))))
+ (fun default : option (defaults.expr (type.base T)) =>
+ (default0 <- default;
+ Some
+ (reflect_list_cps x0
+ (fun ls : option (list (defaults.expr (type.base T))) =>
+ (ls0 <- ls;
+ Some
+ (Some
+ (base.try_make_transport_cps T T
+ (fun
+ a : option
+ (defaults.expr (type.base T) ->
+ defaults.expr (type.base T)) =>
+ match a with
+ | Some x' =>
+ UnderLets.Base
+ (x' (nth_default default0 ls0 args))
+ | None =>
+ UnderLets.Base
+ (ERROR_BAD_REWRITE_RULE
+ (#(pident.List_nth_default) @ ?? @
+ ??{type.base
+ (pattern.base.type.list ??)} @
+ #(pattern.ident.LiteralNat))
+ (#(ident.List_nth_default)%expr @ x @
+ x0 @ x1)%expr_pat)
+ end))));;;
+ None)));;;
+ None)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.List_nth_default)%expr @ x @ x0 @ x1)%expr_pat)%option
| ident.Z_add =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(args + args0)%Z)%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##((args + args0)%Z)%expr
- | None => UnderLets.Base (x + x0)%expr
- end
- | _ => UnderLets.Base (x + x0)%expr
- end
- | None => UnderLets.Base (x + x0)%expr
- end
- | _ => UnderLets.Base (x + x0)%expr
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (x + x0)%expr)%option
| ident.Z_mul =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(args * args0)%Z)%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##((args * args0)%Z)%expr
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
- | None => UnderLets.Base (x * x0)%expr
- end
- | _ => UnderLets.Base (x * x0)%expr
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (x * x0)%expr)%option
| ident.Z_pow =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(args ^ args0))%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##(args ^ args0)%expr
- | None =>
- UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat
- end
- | None => UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat)%option
| ident.Z_sub =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(args - args0)%Z)%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##((args - args0)%Z)%expr
- | None => UnderLets.Base (x - x0)%expr
- end
- | _ => UnderLets.Base (x - x0)%expr
- end
- | None => UnderLets.Base (x - x0)%expr
- end
- | _ => UnderLets.Base (x - x0)%expr
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (x - x0)%expr)%option
| ident.Z_opp =>
fun x : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args => UnderLets.Base ##((- args)%Z)%expr
- | None => UnderLets.Base (- x)%expr
- end
- | _ => UnderLets.Base (- x)%expr
- end
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ Some (UnderLets.Base (##(- args)%Z)%expr)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (- x)%expr)%option
| ident.Z_div =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(args / args0)%Z)%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##((args / args0)%Z)%expr
- | None => UnderLets.Base (x / x0)%expr
- end
- | _ => UnderLets.Base (x / x0)%expr
- end
- | None => UnderLets.Base (x / x0)%expr
- end
- | _ => UnderLets.Base (x / x0)%expr
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (x / x0)%expr)%option
| ident.Z_modulo =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(args mod args0)%Z)%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##((args mod args0)%Z)%expr
- | None => UnderLets.Base (x mod x0)%expr
- end
- | _ => UnderLets.Base (x mod x0)%expr
- end
- | None => UnderLets.Base (x mod x0)%expr
- end
- | _ => UnderLets.Base (x mod x0)%expr
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (x mod x0)%expr)%option
| ident.Z_log2 =>
fun x : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args => UnderLets.Base ##(Z.log2 args)%expr
- | None => UnderLets.Base (#(ident.Z_log2)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_log2)%expr @ x)%expr_pat
- end
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ Some (UnderLets.Base (##(Z.log2 args))%expr)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_log2)%expr @ x)%expr_pat)%option
| ident.Z_log2_up =>
fun x : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args => UnderLets.Base ##(Z.log2_up args)%expr
- | None => UnderLets.Base (#(ident.Z_log2_up)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_log2_up)%expr @ x)%expr_pat
- end
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ Some (UnderLets.Base (##(Z.log2_up args))%expr)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_log2_up)%expr @ x)%expr_pat)%option
| ident.Z_eqb =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(args =? args0))%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##(args =? args0)%expr
- | None =>
- UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat
- end
- | None => UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat)%option
| ident.Z_leb =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(args <=? args0))%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##(args <=? args0)%expr
- | None =>
- UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat
- end
- | None => UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat)%option
| ident.Z_geb =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(args >=? args0))%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##(args >=? args0)%expr
- | None =>
- UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat
- end
- | None => UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat)%option
| ident.Z_of_nat =>
fun x : defaults.expr (type.base base.type.nat) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option nat) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun _ : Z => None
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun v0 : nat => Some v0
- end v
- | _ => None
- end
- with
- | Some args => UnderLets.Base ##(Z.of_nat args)%expr
- | None => UnderLets.Base (#(ident.Z_of_nat)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_of_nat)%expr @ x)%expr_pat
- end
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralNat;
+ Some (UnderLets.Base (##(Z.of_nat args))%expr)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_of_nat)%expr @ x)%expr_pat)%option
| ident.Z_to_nat =>
fun x : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args => UnderLets.Base ##(Z.to_nat args)%expr
- | None => UnderLets.Base (#(ident.Z_to_nat)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_to_nat)%expr @ x)%expr_pat
- end
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ Some (UnderLets.Base (##(Z.to_nat args))%expr)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_to_nat)%expr @ x)%expr_pat)%option
| ident.Z_shiftr =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(Z.shiftr args args0))%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##(Z.shiftr args args0)%expr
- | None => UnderLets.Base (x >> x0)%expr
- end
- | _ => UnderLets.Base (x >> x0)%expr
- end
- | None => UnderLets.Base (x >> x0)%expr
- end
- | _ => UnderLets.Base (x >> x0)%expr
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (x >> x0)%expr)%option
| ident.Z_shiftl =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(Z.shiftl args args0))%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##(Z.shiftl args args0)%expr
- | None => UnderLets.Base (x << x0)%expr
- end
- | _ => UnderLets.Base (x << x0)%expr
- end
- | None => UnderLets.Base (x << x0)%expr
- end
- | _ => UnderLets.Base (x << x0)%expr
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (x << x0)%expr)%option
| ident.Z_land =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(Z.land args args0))%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##(Z.land args args0)%expr
- | None => UnderLets.Base (x &' x0)%expr
- end
- | _ => UnderLets.Base (x &' x0)%expr
- end
- | None => UnderLets.Base (x &' x0)%expr
- end
- | _ => UnderLets.Base (x &' x0)%expr
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (x &' x0)%expr)%option
| ident.Z_lor =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(Z.lor args args0))%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 => UnderLets.Base ##(Z.lor args args0)%expr
- | None => UnderLets.Base (x || x0)%expr
- end
- | _ => UnderLets.Base (x || x0)%expr
- end
- | None => UnderLets.Base (x || x0)%expr
- end
- | _ => UnderLets.Base (x || x0)%expr
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (x || x0)%expr)%option
| ident.Z_bneg =>
fun x : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args => UnderLets.Base ##(Definitions.Z.bneg args)%expr
- | None => UnderLets.Base (#(ident.Z_bneg)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_bneg)%expr @ x)%expr_pat
- end
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ Some (UnderLets.Base (##(Definitions.Z.bneg args))%expr)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_bneg)%expr @ x)%expr_pat)%option
| ident.Z_lnot_modulo =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(Definitions.Z.lnot_modulo args args0))%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- UnderLets.Base
- ##(Definitions.Z.lnot_modulo args args0)%expr
- | None =>
- UnderLets.Base
- (#(ident.Z_lnot_modulo)%expr @ x @ x0)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_lnot_modulo)%expr @ x @ x0)%expr_pat
- end
- | None =>
- UnderLets.Base (#(ident.Z_lnot_modulo)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_lnot_modulo)%expr @ x @ x0)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_lnot_modulo)%expr @ x @ x0)%expr_pat)%option
| ident.Z_mul_split =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (let
+ '(a, b) := Definitions.Z.mul_split args args0 args1
+ in ((##a)%expr, (##b)%expr)%expr_pat))
+ | _ => None
+ end
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- (let
- '(a, b) :=
- Definitions.Z.mul_split args args0 args1 in
- (##(a)%expr, ##(b)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat)%option
| ident.Z_add_get_carry =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (let
+ '(a, b) :=
+ Definitions.Z.add_get_carry_full args args0 args1 in
+ ((##a)%expr, (##b)%expr)%expr_pat))
+ | _ => None
+ end
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- (let
- '(a, b) :=
- Definitions.Z.add_get_carry_full args args0
- args1 in (##(a)%expr, ##(b)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat)%option
| ident.Z_add_with_carry =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(Definitions.Z.add_with_carry args args0 args1))%expr)
+ | _ => None
+ end
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- ##(Definitions.Z.add_with_carry args args0
- args1)%expr
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat)%option
| ident.Z_add_with_get_carry =>
fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ match x2 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (let
+ '(a, b) :=
+ Definitions.Z.add_with_get_carry_full args args0
+ args1 args2 in
+ ((##a)%expr, (##b)%expr)%expr_pat))
| _ => None
end
- with
- | Some args0 =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- UnderLets.Base
- (let
- '(a, b) :=
- Definitions.Z.add_with_get_carry_full
- args args0 args1 args2 in
- (##(a)%expr, ##(b)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
+ | _ => None
+ end
+ | _ => None
+ end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat)%option
| ident.Z_sub_get_borrow =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (let
+ '(a, b) :=
+ Definitions.Z.sub_get_borrow_full args args0 args1
+ in ((##a)%expr, (##b)%expr)%expr_pat))
+ | _ => None
+ end
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- (let
- '(a, b) :=
- Definitions.Z.sub_get_borrow_full args args0
- args1 in (##(a)%expr, ##(b)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat)%option
| ident.Z_sub_with_get_borrow =>
fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ match x2 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (let
+ '(a, b) :=
+ Definitions.Z.sub_with_get_borrow_full args
+ args0 args1 args2 in
+ ((##a)%expr, (##b)%expr)%expr_pat))
| _ => None
end
- with
- | Some args0 =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- UnderLets.Base
- (let
- '(a, b) :=
- Definitions.Z.sub_with_get_borrow_full
- args args0 args1 args2 in
- (##(a)%expr, ##(b)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @
- x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @
- x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @
- x1 @ x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
+ | _ => None
+ end
+ | _ => None
+ end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat)%option
| ident.Z_zselect =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(Definitions.Z.zselect args args0 args1))%expr)
+ | _ => None
+ end
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- ##(Definitions.Z.zselect args args0 args1)%expr
- | None =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat)%option
| ident.Z_add_modulo =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(Definitions.Z.add_modulo args args0 args1))%expr)
+ | _ => None
+ end
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- ##(Definitions.Z.add_modulo args args0 args1)%expr
- | None =>
- UnderLets.Base
- (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat)%option
| ident.Z_rshi =>
fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ match x2 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(Definitions.Z.rshi args args0 args1 args2))%expr)
| _ => None
end
- with
- | Some args0 =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- UnderLets.Base
- ##(Definitions.Z.rshi args args0 args1
- args2)%expr
- | None =>
- UnderLets.Base
- (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @
- x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | None =>
- UnderLets.Base (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- end
+ | _ => None
+ end
+ | _ => None
+ end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat)%option
| ident.Z_cc_m =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ Some (UnderLets.Base (##(Definitions.Z.cc_m args args0))%expr)
| _ => None
end
- with
- | Some args =>
- match x0 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- UnderLets.Base ##(Definitions.Z.cc_m args args0)%expr
- | None =>
- UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat
- end
- | None => UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat)%option
| ident.Z_cast range =>
fun x : defaults.expr (type.base base.type.Z) =>
- match x with
- | #(idc)%expr_pat =>
- match
- match idc with
- | @ident.Literal t0 v =>
- match t0 as t1 return (base.base_interp t1 -> option Z) with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args =>
- UnderLets.Base
- ##(ident.cast ident.cast_outside_of_range range args)%expr
- | None => UnderLets.Base (#(ident.Z_cast range)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_cast range)%expr @ x)%expr_pat
- end
+ ((match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- pattern.ident.invert_bind_args idc pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(ident.cast ident.cast_outside_of_range range args))%expr)
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_cast range)%expr @ x)%expr_pat)%option
| ident.Z_cast2 range =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- match x with
- | (#(idc) @ x1 @ x0)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
+ ((match x with
+ | (@expr.Ident _ _ _ t idc @ x1 @ x0)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (let
+ '(a, b) :=
+ (let (r1, r2) := range in
+ fun '(x2, x3) =>
+ (ident.cast ident.cast_outside_of_range r1 x2,
+ ident.cast ident.cast_outside_of_range r2 x3))
+ (args0, args1) in
+ ((##a)%expr, (##b)%expr)%expr_pat))
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- (let
- '(a, b) :=
- (let (r1, r2) := range in
- fun '(x2, x3) =>
- (ident.cast ident.cast_outside_of_range r1
- x2,
- ident.cast ident.cast_outside_of_range r2 x3))
- (args0, args1) in
- (##(a)%expr, ##(b)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- end
- | None => UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat)%option
| ident.fancy_add log2wordmax imm =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- match x with
- | (#(idc) @ x1 @ x0)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
+ ((match x with
+ | (@expr.Ident _ _ _ t idc @ x1 @ x0)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (let
+ '(a, b) :=
+ ident.fancy.interp
+ (invert_Some
+ (ident.to_fancy
+ (ident.fancy_add log2wordmax imm)))
+ (args0, args1) in
+ ((##a)%expr, (##b)%expr)%expr_pat))
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- (let
- '(a, b) :=
- ident.fancy.interp
- (invert_Some
- (ident.to_fancy
- (ident.fancy_add log2wordmax imm)))
- (args0, args1) in
- (##(a)%expr, ##(b)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat)%option
| ident.fancy_addc log2wordmax imm =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- match x with
- | (#(idc) @ x1 @ x0)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
- | _ => None
- end
- with
- | Some _ =>
- match x1 with
- | (#(idc0) @ x3 @ x2)%expr_pat =>
- match
- match idc0 with
- | @ident.pair A B => Some (A, B)
+ ((match x with
+ | (@expr.Ident _ _ _ t idc @ x1 @ x0)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x1 with
+ | (@expr.Ident _ _ _ t0 idc0 @ x3 @ x2)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.pair;
+ match x3 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ match x2 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t3 idc3 =>
+ args3 <- pattern.ident.invert_bind_args idc3
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (let
+ '(a, b) :=
+ ident.fancy.interp
+ (invert_Some
+ (ident.to_fancy
+ (ident.fancy_addc log2wordmax imm)))
+ (args1, args2, args3) in
+ ((##a)%expr, (##b)%expr)%expr_pat))
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match x0 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args3 =>
- UnderLets.Base
- (let
- '(a, b) :=
- ident.fancy.interp
- (invert_Some
- (ident.to_fancy
- (ident.fancy_addc
- log2wordmax imm)))
- (args1, args2, args3) in
- (##(a)%expr, ##(b)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.fancy_addc log2wordmax
- imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_addc log2wordmax imm)%expr @
- x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_addc log2wordmax imm)%expr @
- x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_addc log2wordmax imm)%expr @
- x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
- end
+ | _ => None
+ end
+ | _ => None
+ end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat)%option
| ident.fancy_sub log2wordmax imm =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- match x with
- | (#(idc) @ x1 @ x0)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
+ ((match x with
+ | (@expr.Ident _ _ _ t idc @ x1 @ x0)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (let
+ '(a, b) :=
+ ident.fancy.interp
+ (invert_Some
+ (ident.to_fancy
+ (ident.fancy_sub log2wordmax imm)))
+ (args0, args1) in
+ ((##a)%expr, (##b)%expr)%expr_pat))
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- (let
- '(a, b) :=
- ident.fancy.interp
- (invert_Some
- (ident.to_fancy
- (ident.fancy_sub log2wordmax imm)))
- (args0, args1) in
- (##(a)%expr, ##(b)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat)%option
| ident.fancy_subb log2wordmax imm =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- match x with
- | (#(idc) @ x1 @ x0)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
- | _ => None
- end
- with
- | Some _ =>
- match x1 with
- | (#(idc0) @ x3 @ x2)%expr_pat =>
- match
- match idc0 with
- | @ident.pair A B => Some (A, B)
+ ((match x with
+ | (@expr.Ident _ _ _ t idc @ x1 @ x0)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x1 with
+ | (@expr.Ident _ _ _ t0 idc0 @ x3 @ x2)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.pair;
+ match x3 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ match x2 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t3 idc3 =>
+ args3 <- pattern.ident.invert_bind_args idc3
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (let
+ '(a, b) :=
+ ident.fancy.interp
+ (invert_Some
+ (ident.to_fancy
+ (ident.fancy_subb log2wordmax imm)))
+ (args1, args2, args3) in
+ ((##a)%expr, (##b)%expr)%expr_pat))
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match x0 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args3 =>
- UnderLets.Base
- (let
- '(a, b) :=
- ident.fancy.interp
- (invert_Some
- (ident.to_fancy
- (ident.fancy_subb
- log2wordmax imm)))
- (args1, args2, args3) in
- (##(a)%expr, ##(b)%expr)%expr_pat)
- | None =>
- UnderLets.Base
- (#(ident.fancy_subb log2wordmax
- imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_subb log2wordmax imm)%expr @
- x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_subb log2wordmax imm)%expr @
- x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_subb log2wordmax imm)%expr @
- x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
- end
+ | _ => None
+ end
+ | _ => None
+ end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat)%option
| ident.fancy_mulll log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- match x with
- | (#(idc) @ x1 @ x0)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
+ ((match x with
+ | (@expr.Ident _ _ _ t idc @ x1 @ x0)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(ident.fancy.interp
+ (invert_Some
+ (ident.to_fancy
+ (ident.fancy_mulll log2wordmax)))
+ (args0, args1)%core))%expr)
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- ##(ident.fancy.interp
- (invert_Some
- (ident.to_fancy
- (ident.fancy_mulll log2wordmax)))
- (args0, args1)%core)%expr
- | None =>
- UnderLets.Base
- (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat)%option
| ident.fancy_mullh log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- match x with
- | (#(idc) @ x1 @ x0)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
+ ((match x with
+ | (@expr.Ident _ _ _ t idc @ x1 @ x0)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(ident.fancy.interp
+ (invert_Some
+ (ident.to_fancy
+ (ident.fancy_mullh log2wordmax)))
+ (args0, args1)%core))%expr)
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- ##(ident.fancy.interp
- (invert_Some
- (ident.to_fancy
- (ident.fancy_mullh log2wordmax)))
- (args0, args1)%core)%expr
- | None =>
- UnderLets.Base
- (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat)%option
| ident.fancy_mulhl log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- match x with
- | (#(idc) @ x1 @ x0)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
+ ((match x with
+ | (@expr.Ident _ _ _ t idc @ x1 @ x0)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(ident.fancy.interp
+ (invert_Some
+ (ident.to_fancy
+ (ident.fancy_mulhl log2wordmax)))
+ (args0, args1)%core))%expr)
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- ##(ident.fancy.interp
- (invert_Some
- (ident.to_fancy
- (ident.fancy_mulhl log2wordmax)))
- (args0, args1)%core)%expr
- | None =>
- UnderLets.Base
- (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat)%option
| ident.fancy_mulhh log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- match x with
- | (#(idc) @ x1 @ x0)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
+ ((match x with
+ | (@expr.Ident _ _ _ t idc @ x1 @ x0)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x1 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(ident.fancy.interp
+ (invert_Some
+ (ident.to_fancy
+ (ident.fancy_mulhh log2wordmax)))
+ (args0, args1)%core))%expr)
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x1 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x0 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- ##(ident.fancy.interp
- (invert_Some
- (ident.to_fancy
- (ident.fancy_mulhh log2wordmax)))
- (args0, args1)%core)%expr
- | None =>
- UnderLets.Base
- (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat)%option
| ident.fancy_rshi log2wordmax x =>
fun x0 : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- match x0 with
- | (#(idc) @ x2 @ x1)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
+ ((match x0 with
+ | (@expr.Ident _ _ _ t idc @ x2 @ x1)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x2 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args0 <- pattern.ident.invert_bind_args idc0 pident.LiteralZ;
+ match x1 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(ident.fancy.interp
+ (invert_Some
+ (ident.to_fancy
+ (ident.fancy_rshi log2wordmax x)))
+ (args0, args1)%core))%expr)
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match idc0 with
- | @ident.Literal t1 v =>
- match
- t1 as t2 return (base.base_interp t2 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args0 =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- UnderLets.Base
- ##(ident.fancy.interp
- (invert_Some
- (ident.to_fancy
- (ident.fancy_rshi log2wordmax x)))
- (args0, args1)%core)%expr
- | None =>
- UnderLets.Base
- (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
- end
- | _ =>
- UnderLets.Base (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
- end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat)%option
| ident.fancy_selc =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- match x with
- | (#(idc) @ x1 @ x0)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
- | _ => None
- end
- with
- | Some _ =>
- match x1 with
- | (#(idc0) @ x3 @ x2)%expr_pat =>
- match
- match idc0 with
- | @ident.pair A B => Some (A, B)
+ ((match x with
+ | (@expr.Ident _ _ _ t idc @ x1 @ x0)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x1 with
+ | (@expr.Ident _ _ _ t0 idc0 @ x3 @ x2)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.pair;
+ match x3 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ match x2 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t3 idc3 =>
+ args3 <- pattern.ident.invert_bind_args idc3
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(ident.fancy.interp
+ (invert_Some
+ (ident.to_fancy ident.fancy_selc))
+ (args1, args2, args3)%core))%expr)
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match x0 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args3 =>
- UnderLets.Base
- ##(ident.fancy.interp
- (invert_Some
- (ident.to_fancy
- ident.fancy_selc))
- (args1, args2, args3)%core)%expr
- | None =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat
- end
- | None => UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat
- end
+ | _ => None
+ end
+ | _ => None
+ end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat)%option
| ident.fancy_selm log2wordmax =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- match x with
- | (#(idc) @ x1 @ x0)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
- | _ => None
- end
- with
- | Some _ =>
- match x1 with
- | (#(idc0) @ x3 @ x2)%expr_pat =>
- match
- match idc0 with
- | @ident.pair A B => Some (A, B)
+ ((match x with
+ | (@expr.Ident _ _ _ t idc @ x1 @ x0)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x1 with
+ | (@expr.Ident _ _ _ t0 idc0 @ x3 @ x2)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.pair;
+ match x3 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ match x2 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t3 idc3 =>
+ args3 <- pattern.ident.invert_bind_args idc3
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(ident.fancy.interp
+ (invert_Some
+ (ident.to_fancy
+ (ident.fancy_selm log2wordmax)))
+ (args1, args2, args3)%core))%expr)
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match x0 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args3 =>
- UnderLets.Base
- ##(ident.fancy.interp
- (invert_Some
- (ident.to_fancy
- (ident.fancy_selm
- log2wordmax)))
- (args1, args2, args3)%core)%expr
- | None =>
- UnderLets.Base
- (#(ident.fancy_selm log2wordmax)%expr @
- x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_selm log2wordmax)%expr @
- x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_selm log2wordmax)%expr @
- x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
- end
+ | _ => None
+ end
+ | _ => None
+ end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat)%option
| ident.fancy_sell =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- match x with
- | (#(idc) @ x1 @ x0)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
- | _ => None
- end
- with
- | Some _ =>
- match x1 with
- | (#(idc0) @ x3 @ x2)%expr_pat =>
- match
- match idc0 with
- | @ident.pair A B => Some (A, B)
+ ((match x with
+ | (@expr.Ident _ _ _ t idc @ x1 @ x0)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x1 with
+ | (@expr.Ident _ _ _ t0 idc0 @ x3 @ x2)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.pair;
+ match x3 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ match x2 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t3 idc3 =>
+ args3 <- pattern.ident.invert_bind_args idc3
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(ident.fancy.interp
+ (invert_Some
+ (ident.to_fancy ident.fancy_sell))
+ (args1, args2, args3)%core))%expr)
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match x0 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args3 =>
- UnderLets.Base
- ##(ident.fancy.interp
- (invert_Some
- (ident.to_fancy
- ident.fancy_sell))
- (args1, args2, args3)%core)%expr
- | None =>
- UnderLets.Base
- (#(ident.fancy_sell)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_sell)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_sell)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_sell)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_sell)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_sell)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat
- end
- | None => UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat
- end
+ | _ => None
+ end
+ | _ => None
+ end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat)%option
| ident.fancy_addm =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- match x with
- | (#(idc) @ x1 @ x0)%expr_pat =>
- match
- match idc with
- | @ident.pair A B => Some (A, B)
- | _ => None
- end
- with
- | Some _ =>
- match x1 with
- | (#(idc0) @ x3 @ x2)%expr_pat =>
- match
- match idc0 with
- | @ident.pair A B => Some (A, B)
+ ((match x with
+ | (@expr.Ident _ _ _ t idc @ x1 @ x0)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc pident.pair;
+ match x1 with
+ | (@expr.Ident _ _ _ t0 idc0 @ x3 @ x2)%expr_pat =>
+ _ <- pattern.ident.invert_bind_args idc0 pident.pair;
+ match x3 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ args1 <- pattern.ident.invert_bind_args idc1
+ pident.LiteralZ;
+ match x2 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ args2 <- pattern.ident.invert_bind_args idc2
+ pident.LiteralZ;
+ match x0 with
+ | @expr.Ident _ _ _ t3 idc3 =>
+ args3 <- pattern.ident.invert_bind_args idc3
+ pident.LiteralZ;
+ Some
+ (UnderLets.Base
+ (##(ident.fancy.interp
+ (invert_Some
+ (ident.to_fancy ident.fancy_addm))
+ (args1, args2, args3)%core))%expr)
+ | _ => None
+ end
| _ => None
end
- with
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal t2 v =>
- match
- t2 as t3
- return (base.base_interp t3 -> option Z)
- with
- | base.type.unit => fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool => fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match x2 with
- | #(idc2)%expr_pat =>
- match
- match idc2 with
- | @ident.Literal t3 v =>
- match
- t3 as t4
- return
- (base.base_interp t4 -> option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z => fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat => fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args2 =>
- match x0 with
- | #(idc3)%expr_pat =>
- match
- match idc3 with
- | @ident.Literal t4 v =>
- match
- t4 as t5
- return
- (base.base_interp t5 ->
- option Z)
- with
- | base.type.unit =>
- fun _ : unit => None
- | base.type.Z =>
- fun v0 : Z => Some v0
- | base.type.bool =>
- fun _ : bool => None
- | base.type.nat =>
- fun _ : nat => None
- end v
- | _ => None
- end
- with
- | Some args3 =>
- UnderLets.Base
- ##(ident.fancy.interp
- (invert_Some
- (ident.to_fancy
- ident.fancy_addm))
- (args1, args2, args3)%core)%expr
- | None =>
- UnderLets.Base
- (#(ident.fancy_addm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_addm)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_addm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_addm)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base
- (#(ident.fancy_addm)%expr @ x)%expr_pat
- end
- | _ =>
- UnderLets.Base
- (#(ident.fancy_addm)%expr @ x)%expr_pat
- end
- | None =>
- UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat
- end
- | None => UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat
- end
- | _ => UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat
- end
+ | _ => None
+ end
+ | _ => None
+ end
+ | _ => None
+ end;;
+ None);;;
+ UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat)%option
end
- : Compile.value' true t
+ : value' true t