aboutsummaryrefslogtreecommitdiff
path: root/src/Experiments/NewPipeline
diff options
context:
space:
mode:
Diffstat (limited to 'src/Experiments/NewPipeline')
-rw-r--r--src/Experiments/NewPipeline/Rewriter.v134
-rw-r--r--src/Experiments/NewPipeline/arith_rewrite_head.out19703
-rw-r--r--src/Experiments/NewPipeline/fancy_rewrite_head.out34598
-rw-r--r--src/Experiments/NewPipeline/nbe_rewrite_head.out1705
4 files changed, 48688 insertions, 7452 deletions
diff --git a/src/Experiments/NewPipeline/Rewriter.v b/src/Experiments/NewPipeline/Rewriter.v
index dd898bc3e..f9f6e22ac 100644
--- a/src/Experiments/NewPipeline/Rewriter.v
+++ b/src/Experiments/NewPipeline/Rewriter.v
@@ -8,7 +8,6 @@ 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.
@@ -426,49 +425,42 @@ Module Compilers.
=> match ctx with
| nil => cont None ctx None
| ctx0 :: ctx'
- => let default := fun 'tt => @eval_decision_tree T ctx default_case cont in
- let bind_default_in f
- := match default_case with
- | Failure => f default
- | _ => (dlet default := default in f default)
- end in
- bind_default_in
- (fun default
- => 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
+ => 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 (f :: x :: ctx') app_case
+ T ctx' icase
(fun k ctx''
- => match ctx'' with
- | f' :: x' :: ctx'''
- => cont k (rApp f' x' alt :: ctx''')
- | _ => cont None ctx
- end)
+ => cont k (rIdent (pident_to_typed pidc args) alt :: ctx''))
| None => default tt
- end
- | rExpr t e
- | rValue t e
- => default tt
- end))
+ 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)
end
| Swap i d'
=> match swap_list 0 i ctx with
@@ -507,8 +499,7 @@ Module Compilers.
(rew : rewrite_rulesT)
(e : rawexpr)
: UnderLets (expr (type_of_rawexpr e))
- := dlet default := UnderLets.Base (expr_of_rawexpr e) in
- eval_decision_tree
+ := eval_decision_tree
(e::nil) d
(fun k ctx default_on_rewrite_failure
=> match k, ctx return UnderLets (expr (type_of_rawexpr e)) with
@@ -532,18 +523,18 @@ Module Compilers.
=> match fv', default_on_rewrite_failure with
| Some fv'', _ => UnderLets.Base fv''
| None, Some default => default tt
- | None, None => default
+ | None, None => UnderLets.Base (expr_of_rawexpr e)
end))%under_lets
| None => match default_on_rewrite_failure with
| Some default => default tt
- | None => default
+ | None => UnderLets.Base (expr_of_rawexpr e)
end
end)
- | None => default
+ | None => UnderLets.Base (expr_of_rawexpr e)
end)
- | None => default
+ | None => UnderLets.Base (expr_of_rawexpr e)
end
- | _, _ => default
+ | _, _ => UnderLets.Base (expr_of_rawexpr e)
end).
Local Notation enumerate ls
@@ -567,20 +558,12 @@ 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 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).
+ := filter (fun '(_, p) => match p with
+ | pattern.Wildcard _::_ => true
+ | _ => false
+ end)
+ p.
Fixpoint get_unique_pattern_ident' (p : list (nat * list pattern)) (so_far : list pident) : list pident
:= match p with
@@ -609,23 +592,25 @@ Module Compilers.
end)
p.
- Definition filter_pattern_app (p : nat * list pattern) : option (nat * list pattern)
+ Definition refine_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 filter_pattern_pident (pidc : pident) (p : nat * list pattern) : option (nat * list pattern)
+ Definition refine_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
@@ -643,14 +628,13 @@ Module Compilers.
=> (onfailure <- compile_rewrites ps;
Some (TryLeaf n1 onfailure))
| Some Datatypes.O
- => let '(pattern_matrix, default_pattern_matrix) := split_at_first_pattern_wildcard pattern_matrix in
- default_case <- compile_rewrites default_pattern_matrix;
+ => default_case <- compile_rewrites (filter_pattern_wildcard pattern_matrix);
app_case <- (if contains_pattern_app pattern_matrix
- then option_map Some (compile_rewrites (Option.List.map filter_pattern_app pattern_matrix))
+ then option_map Some (compile_rewrites (Option.List.map refine_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 (filter_pattern_pident pidc) pattern_matrix)))
+ (fun pidc => option_map (pair pidc) (compile_rewrites (Option.List.map (refine_pattern_pident pidc) pattern_matrix)))
pidcs in
Some (Switch icases app_case default_case)
| Some i
@@ -1740,7 +1724,6 @@ 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
- Let_In
UnderLets.splice UnderLets.to_expr
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
@@ -1804,9 +1787,6 @@ 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.
@@ -1820,7 +1800,6 @@ 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
- Let_In
UnderLets.splice UnderLets.to_expr
Compile.reflect UnderLets.reify_and_let_binds_base_cps Compile.reify Compile.reify_and_let_binds_cps
Compile.value'
@@ -1885,9 +1864,6 @@ 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.
@@ -1901,7 +1877,6 @@ 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
- Let_In
UnderLets.splice UnderLets.to_expr
Compile.reflect UnderLets.reify_and_let_binds_base_cps Compile.reify Compile.reify_and_let_binds_cps
Compile.value'
@@ -1966,9 +1941,6 @@ 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/arith_rewrite_head.out b/src/Experiments/NewPipeline/arith_rewrite_head.out
index 994eaa810..23b5205d2 100644
--- a/src/Experiments/NewPipeline/arith_rewrite_head.out
+++ b/src/Experiments/NewPipeline/arith_rewrite_head.out
@@ -8,57 +8,39 @@ 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 => llet default := UnderLets.Base ##(v0)%expr in
- default
- | base.type.Z =>
- fun v0 : Z => llet default := UnderLets.Base ##(v0)%expr in
- default
- | base.type.bool =>
- fun v0 : bool => llet default := UnderLets.Base ##(v0)%expr in
- default
- | base.type.nat =>
- fun v0 : nat => llet default := UnderLets.Base ##(v0)%expr in
- default
+ | 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) =>
- llet default := UnderLets.Base (#(ident.Nat_succ)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Nat_succ)%expr @ x)%expr_pat
| ident.Nat_pred =>
fun x : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_pred)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Nat_pred)%expr @ x)%expr_pat
| ident.Nat_max =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat
| ident.Nat_mul =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat
| ident.Nat_add =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat
| ident.Nat_sub =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat in
- default
-| @ident.nil t => llet default := UnderLets.Base []%expr_pat in
- default
+ UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat
+| @ident.nil t => UnderLets.Base []%expr_pat
| @ident.cons t =>
fun (x : defaults.expr (type.base t))
(x0 : defaults.expr (type.base (base.type.list t))) =>
- llet default := UnderLets.Base (x :: x0)%expr_pat in
- default
+ UnderLets.Base (x :: x0)%expr_pat
| @ident.pair A B =>
fun (x : defaults.expr (type.base A)) (x0 : defaults.expr (type.base B))
- => llet default := UnderLets.Base (x, x0)%expr_pat in
- default
+ => UnderLets.Base (x, x0)%expr_pat
| @ident.fst A B =>
fun x : defaults.expr (type.base (A * B)%etype) =>
- llet default := UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat in
match x with
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x1) x0 =>
match
@@ -97,30 +79,32 @@ match idc in (ident t) return (Compile.value' true t) with
defaults.expr (type.base A)) =>
match a with
| Some x' => UnderLets.Base (x' v)
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
end)
| (s1 -> d1)%ptype =>
fun _ : Compile.value' false s1 -> Compile.value' true d1
- => default
+ => UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
end (Compile.reflect x0)
| (s1 -> d1)%ptype =>
fun _ : Compile.value' false s1 -> Compile.value' true d1 =>
- default
+ UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
end (Compile.reflect x1)
- | None => default
+ | 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 _ _ _ _ _ _ _) _) _ => default
+ 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 _ _ _ _ _ _ _) _ => default
- | _ => default
+ (@expr.LetIn _ _ _ _ _ _ _) _ =>
+ UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
+ | _ => UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
end
| @ident.snd A B =>
fun x : defaults.expr (type.base (A * B)%etype) =>
- llet default := UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat in
match x with
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x1) x0 =>
match
@@ -159,26 +143,29 @@ match idc in (ident t) return (Compile.value' true t) with
defaults.expr (type.base B)) =>
match a with
| Some x' => UnderLets.Base (x' v0)
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
end)
| (s1 -> d1)%ptype =>
fun _ : Compile.value' false s1 -> Compile.value' true d1
- => default
+ => UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
end (Compile.reflect x0)
| (s1 -> d1)%ptype =>
fun _ : Compile.value' false s1 -> Compile.value' true d1 =>
- default
+ UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
end (Compile.reflect x1)
- | None => default
+ | 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 _ _ _ _ _ _ _) _) _ => default
+ 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 _ _ _ _ _ _ _) _ => default
- | _ => default
+ (@expr.LetIn _ _ _ _ _ _ _) _ =>
+ UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
+ | _ => UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
end
| @ident.prod_rect A B T =>
fun
@@ -187,11 +174,10 @@ 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)) =>
- llet default := 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 in
- default
+ 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
| @ident.bool_rect T =>
fun
(x
@@ -199,13 +185,12 @@ 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)) =>
- llet default := 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 in
- default
+ 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
| @ident.nat_rect P =>
fun
(x : defaults.expr (type.base base.type.unit) ->
@@ -216,15 +201,12 @@ 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)) =>
- llet default := 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 in
- default
+ 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
| @ident.nat_rect_arrow P Q =>
fun
(x : defaults.expr (type.base P) ->
@@ -239,22 +221,18 @@ 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)) =>
- llet default := 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 in
- default
+ 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
| @ident.list_rect A P =>
fun
(x : defaults.expr (type.base base.type.unit) ->
@@ -266,19 +244,13 @@ 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))) =>
- llet default := 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 in
- default
+ 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
| @ident.list_case A P =>
fun
(x : defaults.expr (type.base base.type.unit) ->
@@ -289,89 +261,70 @@ 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))) =>
- llet default := 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 in
- default
+ 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
| @ident.List_length T =>
fun x : defaults.expr (type.base (base.type.list T)) =>
- llet default := UnderLets.Base (#(ident.List_length)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_length)%expr @ x)%expr_pat
| ident.List_seq =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat
| @ident.List_firstn A =>
fun (x : defaults.expr (type.base base.type.nat))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- llet default := UnderLets.Base
- (#(ident.List_firstn)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_firstn)%expr @ x @ x0)%expr_pat
| @ident.List_skipn A =>
fun (x : defaults.expr (type.base base.type.nat))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- llet default := UnderLets.Base
- (#(ident.List_skipn)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_skipn)%expr @ x @ x0)%expr_pat
| @ident.List_repeat A =>
fun (x : defaults.expr (type.base A))
(x0 : defaults.expr (type.base base.type.nat)) =>
- llet default := UnderLets.Base
- (#(ident.List_repeat)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_repeat)%expr @ x @ x0)%expr_pat
| @ident.List_combine A B =>
fun (x : defaults.expr (type.base (base.type.list A)))
(x0 : defaults.expr (type.base (base.type.list B))) =>
- llet default := UnderLets.Base
- (#(ident.List_combine)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_combine)%expr @ x @ x0)%expr_pat
| @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))) =>
- llet default := UnderLets.Base
- (#(ident.List_map)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat in
- default
+ UnderLets.Base
+ (#(ident.List_map)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
| @ident.List_app A =>
fun x x0 : defaults.expr (type.base (base.type.list A)) =>
- llet default := UnderLets.Base (x ++ x0)%expr in
- default
+ UnderLets.Base (x ++ x0)%expr
| @ident.List_rev A =>
fun x : defaults.expr (type.base (base.type.list A)) =>
- llet default := UnderLets.Base (#(ident.List_rev)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_rev)%expr @ x)%expr_pat
| @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))) =>
- llet default := UnderLets.Base
- (#(ident.List_flat_map)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat in
- default
+ UnderLets.Base
+ (#(ident.List_flat_map)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
| @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))) =>
- llet default := UnderLets.Base
- (#(ident.List_partition)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat in
- default
+ UnderLets.Base
+ (#(ident.List_partition)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
| @ident.List_fold_right A B =>
fun
(x : defaults.expr (type.base B) ->
@@ -380,413 +333,511 @@ 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))) =>
- llet default := 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 in
- default
+ 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
| @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))) =>
- llet default := UnderLets.Base
- (#(ident.List_update_nth)%expr @ x @
- (λ x2 : var (type.base T),
- UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat in
- default
+ UnderLets.Base
+ (#(ident.List_update_nth)%expr @ x @
+ (λ x2 : var (type.base T),
+ UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat
| @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)) =>
- llet default := UnderLets.Base
- (#(ident.List_nth_default)%expr @ x @ x0 @ x1)%expr_pat in
- default
+ 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) =>
- llet default := UnderLets.Base (x + x0)%expr in
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- match x0 with
- | @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
- 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 =>
- default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ :
- Compile.value'
- false s0 ->
- Compile.value'
- true d0
- => default
- end
- (Compile.reflect
- x1)
- | None => default
- end
- | @expr.App _ _ _ s _
- ($_)%expr _ |
- @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _
- _) _ | @expr.App _
- _ _ s _
- (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _
- _ _ _) _ => default
- | _ => default
- end in
- match x with
- | @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
- 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
+ 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 t2) ->
+ (type.base t4) ->
defaults.expr
(type.base
base.type.Z))
- =>
- match a with
- | Some x' =>
- UnderLets.Base
- (x0 - x' v)%expr
- | None => default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false
- s0 ->
- Compile.value' true
- d0 => default
- end (Compile.reflect x1)
- | None => default
- end
- | @expr.App _ _ _ s _ ($_)%expr _ |
- @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s _ (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- default
- | _ => default0 tt
- end in
- 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 x with
- | @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
- (-
- (x'0 v0 +
- ##((- args)%Z)))%expr
- else default
- | None => default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value'
- false s0 ->
- Compile.value'
- true d0 =>
- default
- end (Compile.reflect x1)
- | None => default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- default
- end (Compile.reflect x1)
- | None => default
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _
- _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s _ (_ @ _)%expr_pat _ | @expr.App _ _
- _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- default
- | _ => default
- end
- | None => default0 tt
- end
- | @expr.App _ _ _ s _ #(idc)%expr_pat x1 =>
+ =>
+ 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
+ (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
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
+ s as t2
+ return
+ (Compile.value' false t2 ->
+ UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z)))
with
- | Some _ =>
- 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 =>
+ | 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 t2
+ s as t3
return
- (Compile.value' false t2 ->
+ (Compile.value' false t3 ->
UnderLets.UnderLets base.type ident
var
(defaults.expr
(type.base base.type.Z)))
with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
+ | type.base t3 =>
+ fun
+ v0 : defaults.expr (type.base t3)
=>
base.try_make_transport_cps
(fun t1 : base.type =>
- defaults.expr (type.base t1)) t2
+ defaults.expr (type.base t1)) t3
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 >? 0
+ 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
- (##(args0) - x' v)%expr
+ (-
+ (x'0 v0 + ##((- args)%Z)))%expr
else
match
- s as t3
+ s as t4
return
- (Compile.value' false t3 ->
+ (Compile.value' false t4 ->
UnderLets.UnderLets
base.type ident var
(defaults.expr
(type.base
base.type.Z)))
with
- | type.base t3 =>
+ | type.base t4 =>
fun
- v0 : defaults.expr
- (type.base t3)
+ v1 : defaults.expr
+ (type.base t4)
=>
base.try_make_transport_cps
(fun t1 : base.type =>
defaults.expr
- (type.base t1)) t3
+ (type.base t1)) t4
base.type.Z
(UnderLets.UnderLets
base.type ident var
@@ -794,24 +845,21 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a0 : option
+ a1 : option
(defaults.expr
(type.base
- t3) ->
+ t4) ->
defaults.expr
(type.base
base.type.Z))
=>
- match a0 with
- | Some x'0 =>
- if args0 <? 0
- then
- UnderLets.Base
- (-
- (##((- args0)%Z) +
- x'0 v0))%expr
- else default
- | None => default
+ match a1 with
+ | Some x'1 =>
+ UnderLets.Base
+ (x0 - x'1 v1)%expr
+ | None =>
+ UnderLets.Base
+ (x + x0)%expr
end)
| (s0 -> d0)%ptype =>
fun
@@ -819,128 +867,645 @@ match idc in (ident t) return (Compile.value' true t) with
false s0 ->
Compile.value'
true d0 =>
- default
+ UnderLets.Base
+ (x + x0)%expr
end (Compile.reflect x1)
- | None => default
+ | None =>
+ UnderLets.Base (x + x0)%expr
end)
| (s0 -> d0)%ptype =>
fun
_ : Compile.value' false s0 ->
Compile.value' true d0 =>
- default
+ UnderLets.Base (x + x0)%expr
end (Compile.reflect x1)
- | None => default
- end
- | @expr.App _ _ _ s0 _ #(idc0)%expr_pat x2 =>
- 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
+ | 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)))
- 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
- 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' v + x'0 v0))%expr
- | None => default
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 => default
- end (Compile.reflect x1)
- | None => default
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- default
- end (Compile.reflect x2)
- | None => default
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _
- _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s0 _ (_ @ _)%expr_pat _ | @expr.App _ _
- _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- default
- | _ => default
- end
- | None => default
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => default
- | _ => default0 tt
- end in
+ (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
+ (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
@@ -955,1045 +1520,1206 @@ match idc in (ident t) return (Compile.value' true t) with
| _ => None
end
with
- | Some args => if args =? 0 then UnderLets.Base x0 else default
- | None => default0 tt
+ | 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
+ 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 _ =>
+ 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 =? -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
+ 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 (- 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
+ (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
+ (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 _ =>
+ 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 =? -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
+ 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
- | _ => default0 tt
end
-| ident.Z_mul =>
+| 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) =>
- llet default := UnderLets.Base (x * x0)%expr in
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- 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 =?
- 2
- ^ Z.log2
- args) &&
- negb
- (args =? 2)
- then
- UnderLets.Base
- (x0 <<
- ##
- (Z.log2
- args))%expr
- else default
- | None => default
- end
- | _ => default
- end in
- 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 =?
- 2
- ^ Z.log2
- args) &&
- negb
- (args =? 2)
- then
- UnderLets.Base
- (x <<
- ##
- (Z.log2
- args))%expr
- else default
- | None =>
- default0 tt
- 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
- 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 =>
- default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ :
- Compile.value'
- false s0 ->
- Compile.value'
- true d0
- => default
- end
- (Compile.reflect
- x1)
- | None => default
- end
- | @expr.App _ _ _ s _
- ($_)%expr _ |
- @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _
- _) _ | @expr.App _
- _ _ s _
- (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _
- _ _ _) _ => default
- | _ => default0 tt
- end in
- match x with
- | @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
- 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 =>
- default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ :
- Compile.value'
- false s0 ->
- Compile.value'
- true d0
- => default
- end
- (Compile.reflect
- x1)
- | None => default
- end
- | @expr.App _ _ _ s _
- ($_)%expr _ |
- @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _
- _) _ | @expr.App _
- _ _ s _
- (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _
- _ _ _) _ => default
- | _ => default0 tt
- end in
- 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 *
- ##
- ((- args)%Z)))%expr
- else default
- | None =>
- default0 tt
- end
- | @expr.App _ _ _ s _
- #(idc)%expr_pat
- x1 =>
- match
- match idc with
- | ident.Z_opp =>
- Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x with
- | @expr.App _
- _ _ s0 _
- #(idc0)%expr_pat
- x2 =>
- 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)))
- 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
- 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' v *
- x'0 v0)%expr
- | None =>
- default
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- default
- end
- (Compile.reflect
- x1)
- | None =>
- default
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- default
- end
- (Compile.reflect
- x2)
- | None =>
- default
- end
- | @expr.App _
- _ _ s0 _
- ($_)%expr
- _ |
- @expr.App _
- _ _ s0 _
- (@expr.Abs _
- _ _ _ _ _)
- _ |
- @expr.App _
- _ _ s0 _
- (_ @ _)%expr_pat
- _ |
- @expr.App _
- _ _ s0 _
- (@expr.LetIn
- _ _ _ _ _ _
- _) _ =>
- default
- | _ => default
- end
- | None => default
- end
- | @expr.App _ _ _ s _
- ($_)%expr _ |
- @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _
- _) _ | @expr.App _
- _ _ s _
- (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _
- _ _ _) _ => default
- | _ => default0 tt
- end in
- 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
- (-
- (##
- ((- args)%Z) *
- x0))%expr
- else default
- | None =>
- default0 tt
- end
- | _ => default0 tt
- end in
- 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)%expr
- else default
- | None =>
- default0 tt
- end
- | _ => default0 tt
- end in
- 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 =? -1
- then
- UnderLets.Base
- (- x0)%expr
- else default
- | None =>
- default0 tt
- end
- | _ => default0 tt
- end in
- 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
- match x with
- | @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 =?
- -1
- then
- UnderLets.Base
- (x' v)
- else
- default
- | None =>
- default
- end)
- | (s0 ->
- d0)%ptype =>
- fun
- _ :
- Compile.value'
- false s0 ->
- Compile.value'
- true d0
- =>
- default
- end
- (Compile.reflect
- x1)
- | None =>
- default
- end
- | @expr.App _
- _ _ s _
- ($_)%expr
- _ |
- @expr.App _
- _ _ s _
- (@expr.Abs
- _ _ _ _ _
- _) _ |
- @expr.App _
- _ _ s _
- (_ @ _)%expr_pat
- _ |
- @expr.App _
- _ _ s _
- (@expr.LetIn
- _ _ _ _ _
- _ _) _ =>
- default
- | _ =>
- default
- end
- | None =>
- default0 tt
- end
- | @expr.App _ _ _ s _
- #(idc)%expr_pat
- x1 =>
- match
- match idc with
- | ident.Z_opp =>
- Some tt
- | _ => None
- end
- with
- | Some _ =>
- 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 =>
- 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
- (x' v)
- else
- default
- | None =>
- default
- end)
- | (s0 ->
- d0)%ptype =>
- fun
- _ :
- Compile.value'
- false s0 ->
- Compile.value'
- true d0
- =>
- default
- end
- (Compile.reflect
- x1)
- | None =>
- default
- end
- | _ => default
- end
- | None => default
- end
- | @expr.App _ _ _ s _
- ($_)%expr _ |
- @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _
- _) _ | @expr.App _
- _ _ s _
- (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _
- _ _ _) _ => default
- | _ => default0 tt
- end in
- 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 =? 1
- then UnderLets.Base x0
- else default
- | None => default0 tt
- end
- | _ => default0 tt
- end in
- 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 default
- | None => default0 tt
- end
- | _ => default0 tt
- end in
match x with
| #(idc)%expr_pat =>
match
@@ -2009,846 +2735,753 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args =>
- if args =? 0 then UnderLets.Base ##(0)%expr else default
- | None => default0 tt
- end
- | _ => default0 tt
- end
-| ident.Z_pow =>
- fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat in
- default
-| ident.Z_sub =>
- fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x - x0)%expr in
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- match x0 with
- | @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
- 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 =>
- default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ :
- Compile.value'
- false s0 ->
- Compile.value'
- true d0
- => default
- end
- (Compile.reflect
- x1)
- | None => default
- end
- | @expr.App _ _ _ s _
- ($_)%expr _ |
- @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _
- _) _ | @expr.App _
- _ _ s _
- (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _
- _ _ _) _ => default
- | _ => default
- end in
- match x with
- | @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
+ 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
- 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
+ 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
- a :
+ (fun
+ a2 :
option
(defaults.expr
(type.base
- t2) ->
+ t5) ->
defaults.expr
(type.base
base.type.Z))
- =>
- match
- a
- with
- | Some
- x' =>
- UnderLets.Base
- (-
- (x' v +
- x0))%expr
- | None =>
- default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ :
- Compile.value'
- false s0 ->
- Compile.value'
- true d0
- => default
- end
- (Compile.reflect
- x1)
- | None => default
- end
- | @expr.App _ _ _ s _
- ($_)%expr _ |
- @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _
- _) _ | @expr.App _
- _ _ s _
- (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _
- _ _ _) _ => default
- | _ => default0 tt
- end in
- 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)
+ a2
with
- | base.type.unit =>
- fun
- _ : unit
- => None
- | base.type.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 +
- ##
- ((- args)%Z))%expr
- else default
- | None =>
- default0 tt
- end
- | @expr.App _ _ _ s _
- #(idc)%expr_pat
- x1 =>
- match
- match idc with
- | ident.Z_opp =>
- Some tt
- | _ => None
- end
- with
- | Some _ =>
- match x with
- | @expr.App _
- _ _ s0 _
- #(idc0)%expr_pat
- x2 =>
- 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)))
- 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
- 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 =>
+ | Some
+ x'2 =>
UnderLets.Base
- (x'0 v0 -
- x' v)%expr
- | None =>
- default
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- default
- end
- (Compile.reflect
- x1)
- | None =>
- default
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- default
- end
- (Compile.reflect
- x2)
- | None =>
- default
- end
- | @expr.App _
- _ _ s0 _
- ($_)%expr
- _ |
- @expr.App _
- _ _ s0 _
- (@expr.Abs _
- _ _ _ _ _)
- _ |
- @expr.App _
- _ _ s0 _
- (_ @ _)%expr_pat
- _ |
- @expr.App _
- _ _ s0 _
- (@expr.LetIn
- _ _ _ _ _ _
- _) _ =>
- default
- | _ => default
- end
- | None => default
- end
- | @expr.App _ _ _ s _
- ($_)%expr _ |
- @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _
- _) _ | @expr.App _
- _ _ s _
- (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _
- _ _ _) _ => default
- | _ => default0 tt
- end in
- 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
- (- (##((- args)%Z) + x0))%expr
- else default
- | None => default0 tt
- end
- | @expr.App _ _ _ s _ #(idc)%expr_pat
- x1 =>
- match
- match idc with
- | ident.Z_opp => Some tt
- | _ => 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 =>
- 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 >? 0
- then
- UnderLets.Base
- (-
- (x' v +
- ##
- ((- args0)%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
- args0 <?
- 0
- then
+ (x +
+ x'2 v2)%expr
+ | None =>
UnderLets.Base
- (##
- ((- args0)%Z) -
- x'0 v0)%expr
- else
- default
- | None =>
- default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ :
- Compile.value'
+ (x - x0)%expr
+ end)
+ | (s0 -> d0)%ptype =>
+ fun
+ _ : Compile.value'
false s0 ->
Compile.value'
true d0
- =>
- default
- end
- (Compile.reflect
- x1)
- | None =>
- default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value'
- false s0 ->
- Compile.value'
- true d0 =>
- default
- end
- (Compile.reflect x1)
- | None => default
- end
- | _ => default
- end
- | None => default
- end
- | @expr.App _ _ _ s _ ($_)%expr _ |
- @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s _ (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ =>
- default
- | _ => default0 tt
- end in
- 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 default
- | None => default0 tt
- end
- | @expr.App _ _ _ s _ #(idc)%expr_pat x1 =>
+ =>
+ 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
+ (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
+ (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
- match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end
+ s as t2
+ return
+ (Compile.value' false t2 ->
+ UnderLets.UnderLets base.type ident var
+ (defaults.expr (type.base base.type.Z)))
with
- | Some _ =>
- 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 =>
+ | 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 t2
+ s as t3
return
- (Compile.value' false t2 ->
+ (Compile.value' false t3 ->
UnderLets.UnderLets base.type ident
var
(defaults.expr
(type.base base.type.Z)))
with
- | type.base t2 =>
- fun v : defaults.expr (type.base t2)
+ | type.base t3 =>
+ fun
+ v0 : defaults.expr (type.base t3)
=>
base.try_make_transport_cps
(fun t1 : base.type =>
- defaults.expr (type.base t1)) t2
+ defaults.expr (type.base t1)) t3
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 >? 0
+ 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
- (##(args0) + x' v)%expr
+ (##((- args)%Z) - x'0 v0)%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
+ 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
- t3) ->
- defaults.expr
+ t4) ->
+ defaults.expr
(type.base
base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
- if args0 <? 0
- then
+ =>
+ match a1 with
+ | Some x'1 =>
UnderLets.Base
- (x'0 v0 -
- ##((- args0)%Z))%expr
- else default
- | None => default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value'
- false s0 ->
- Compile.value'
- true d0 =>
- default
- end (Compile.reflect x1)
- | None => default
+ (-
+ (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 =>
- default
+ UnderLets.Base (x - x0)%expr
end (Compile.reflect x1)
- | None => default
- end
- | _ => default
- end
- | None => default
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => default
- | _ => default0 tt
- end in
- 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 =>
- llet default1 := fun 'tt =>
- if args =? 0
- then UnderLets.Base (- x0)%expr
- else default in
- match x0 with
- | @expr.App _ _ _ s _ #(idc0)%expr_pat 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 idc0 with
+ 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
@@ -2864,6 +3497,156 @@ match idc in (ident t) return (Compile.value' true t) 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
@@ -2873,40 +3656,368 @@ match idc in (ident t) return (Compile.value' true t) with
(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 default
- | None => default
+ | 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 => default
+ Compile.value' true d0 =>
+ UnderLets.Base (x - x0)%expr
end (Compile.reflect x1)
- | None => default
+ | 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 _ _ _ _ _ _ _) _ => default
- | _ => default1 tt
+ | _ => UnderLets.Base (x - x0)%expr
end
- | None => default0 tt
end
- | _ => default0 tt
+ | @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
+ (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_opp =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (- x)%expr in
- llet default0 := fun 'tt =>
- 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 default in
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
@@ -2933,22 +4044,44 @@ match idc in (ident t) return (Compile.value' true t) with
defaults.expr (type.base base.type.Z)) =>
match a with
| Some x' => UnderLets.Base (x' v)
- | None => default
+ | None => UnderLets.Base (- x)%expr
end)
| (s0 -> d0)%ptype =>
fun _ : Compile.value' false s0 -> Compile.value' true d0 =>
- default
+ UnderLets.Base (- x)%expr
end (Compile.reflect x0)
- | None => default
+ | 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 _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => default
- | _ => default0 tt
+ | @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)
+ 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
| ident.Z_div =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x / x0)%expr in
match x0 with
| #(idc)%expr_pat =>
match
@@ -2969,14 +4102,13 @@ match idc in (ident t) return (Compile.value' true t) with
else
if args =? 2 ^ Z.log2 args
then UnderLets.Base (x >> ##(Z.log2 args))%expr
- else default
- | None => default
+ else UnderLets.Base (x / x0)%expr
+ | None => UnderLets.Base (x / x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x / x0)%expr
end
| ident.Z_modulo =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x mod x0)%expr in
match x0 with
| #(idc)%expr_pat =>
match
@@ -2997,76 +4129,40 @@ match idc in (ident t) return (Compile.value' true t) with
else
if args =? 2 ^ Z.log2 args
then UnderLets.Base (x &' ##((args - 1)%Z))%expr
- else default
- | None => default
+ else UnderLets.Base (x mod x0)%expr
+ | None => UnderLets.Base (x mod x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x mod x0)%expr
end
| ident.Z_log2 =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_log2)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_log2)%expr @ x)%expr_pat
| ident.Z_log2_up =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_log2_up)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_log2_up)%expr @ x)%expr_pat
| ident.Z_eqb =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat
| ident.Z_leb =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat
| ident.Z_geb =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat
| ident.Z_of_nat =>
fun x : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Z_of_nat)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_of_nat)%expr @ x)%expr_pat
| ident.Z_to_nat =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_to_nat)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_to_nat)%expr @ x)%expr_pat
| ident.Z_shiftr =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x >> x0)%expr in
- default
+ UnderLets.Base (x >> x0)%expr
| ident.Z_shiftl =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x << x0)%expr in
- default
+ UnderLets.Base (x << x0)%expr
| ident.Z_land =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x &' x0)%expr in
- llet default0 := fun 'tt =>
- 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 default
- | None => default
- end
- | _ => default
- end in
match x0 with
| #(idc)%expr_pat =>
match
@@ -3082,38 +4178,121 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args =>
- if args =? 0 then UnderLets.Base ##(0)%expr else default
- | None => default0 tt
+ 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
- | _ => default0 tt
end
| ident.Z_lor =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x || x0)%expr in
- default
+ UnderLets.Base (x || x0)%expr
| ident.Z_bneg =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_bneg)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_bneg)%expr @ x)%expr_pat
| ident.Z_lnot_modulo =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_lnot_modulo)%expr @ x @ x0)%expr_pat in
- default
+ 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) =>
- llet default := UnderLets.Base
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat in
- llet default0 := fun 'tt =>
- 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) in
match x with
| #(idc)%expr_pat =>
match
@@ -3129,248 +4308,6 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some _ =>
- llet default1 := fun 'tt =>
- llet default1 := fun 'tt =>
- llet default1 := fun 'tt =>
- llet default1 :=
- fun 'tt =>
- llet default1 :=
- fun 'tt =>
- 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 =?
- -1
- then
- UnderLets.Base
- (
- (- x0)%expr,
- ##
- (0)%expr)%expr_pat
- else
- default
- | None =>
- default
- end
- | _ => default
- end in
- 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 =?
- -1
- then
- UnderLets.Base
- (
- (- x1)%expr,
- ##
- (0)%expr)%expr_pat
- else
- default
- | None =>
- default1
- tt
- end
- | _ =>
- default1
- tt
- end in
- 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 =?
- 1
- then
- UnderLets.Base
- (x0,
- ##
- (0)%expr)%expr_pat
- else
- default
- | None =>
- default1
- tt
- end
- | _ =>
- default1
- tt
- end in
- 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 =? 1
- then
- UnderLets.Base
- (x1, ##(0)%expr)%expr_pat
- else default
- | None => default1 tt
- end
- | _ => default1 tt
- end in
- 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 default
- | None => default1 tt
- end
- | _ => default1 tt
- end in
match x0 with
| #(idc0)%expr_pat =>
match
@@ -3390,305 +4327,1058 @@ match idc in (ident t) return (Compile.value' true t) with
| Some args0 =>
if args0 =? 0
then UnderLets.Base (##(0)%expr, ##(0)%expr)%expr_pat
- else default
- | None => default1 tt
- end
- | _ => default1 tt
- end
- | None => default0 tt
- end
- | _ => default0 tt
- end
-| ident.Z_add_get_carry =>
- fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat in
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- 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) in
- 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.Base
- (x0,
- ##(0)%expr)%expr_pat
- else default
- | None =>
- default0 tt
- end
- | _ => default0 tt
- end in
- 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
- (x1,
- ##(0)%expr)%expr_pat
- else default
- | None =>
- default0 tt
- end
- | _ => default0 tt
- end in
- 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 default
- | None =>
- default0 tt
- end
- | _ => default0 tt
- end in
- 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.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 default
- | None => default0 tt
- end
- | _ => default0 tt
- end in
+ else
match x1 with
- | @expr.App _ _ _ s _ #(idc)%expr_pat x2 =>
+ | #(idc1)%expr_pat =>
match
- match idc with
- | ident.Z_opp => Some tt
+ 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 _ =>
- 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 => default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 => default
- end (Compile.reflect x2)
- | None => default
+ | 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.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => default
- | _ => default0 tt
- end in
+ | ($_)%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
+| 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
@@ -3726,28 +5416,867 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base
(#(ident.fst)%expr @ ($v0)%expr,
(- (#(ident.snd)%expr @ $v0)%expr_pat)%expr)%expr_pat)
- | None => default
+ | 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 =>
- default
+ UnderLets.Base
+ (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
end (Compile.reflect x2)
- | None => default
+ | 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
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => default
- | _ => default0 tt
end
| ident.Z_add_with_carry =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat in
- llet default0 := fun 'tt =>
- 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) in
match x with
| #(idc)%expr_pat =>
match
@@ -3763,55 +6292,340 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args =>
- if args =? 0 then UnderLets.Base (x0 + x1)%expr else default
- | None => default0 tt
+ 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
- | _ => default0 tt
+ | ($_)%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) =>
- llet default := UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat in
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- 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) in
- 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 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 idc0 with
- | @ident.Literal t1 v =>
+ match idc2 with
+ | @ident.Literal t3 v =>
match
- t1 as t2
+ t3 as t4
return
- (base.base_interp t2 -> option Z)
+ (base.base_interp t4 -> option Z)
with
| base.type.unit =>
fun _ : unit => None
@@ -3823,13 +6637,55 @@ match idc in (ident t) return (Compile.value' true t) with
| _ => None
end
with
- | Some args0 =>
+ | 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 @ ##(args)%expr @
- ##(args0)%expr)%expr_pat
+ x @ x0 @ x1 @ x2)%expr_pat
(fun
v : var
(type.base
@@ -3837,514 +6693,2494 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)%etype) =>
UnderLets.Base
(#(ident.fst)%expr @ ($v)%expr,
- ##(0)%expr)%expr_pat)
- else default
- | None => default
+ #(ident.snd)%expr @ ($v)%expr)%expr_pat)
end
- | _ => default
+ | ($_)%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
+ (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
- | None => default0 tt
+ | ($_)%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
- | _ => default0 tt
- end in
- 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 =>
- llet default1 := fun 'tt =>
- llet default1 := fun 'tt =>
- llet default1 := fun 'tt =>
- llet default1 :=
- fun 'tt =>
- llet default1 :=
- fun 'tt =>
- if args =? 0
- then
- UnderLets.UnderLet
- (#(ident.Z_add_get_carry)%expr @
- x @ x1 @
- x2)%expr_pat
- (fun
- v :
+ | @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
+ =>
+ UnderLets.Base
(
#
(ident.fst)%expr @
- ($v)%expr,
+ ($v1)%expr,
#
(ident.snd)%expr @
- ($v)%expr)%expr_pat)
- else default in
- 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
- _ =>
- llet default2 :=
- fun 'tt
- =>
- 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
- (args =?
- 0) &&
- (args1 =?
- 0)
- then
- UnderLets.Base
- (x1,
- ##
- (0)%expr)%expr_pat
- else
- default
- | None =>
- default
- end
- | _ =>
- default
- end in
- 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
- (args =?
- 0) &&
- (args1 =?
- 0)
- then
- UnderLets.Base
- (x2,
- ##
- (0)%expr)%expr_pat
- else
- default
- | None =>
- default2
- tt
- end
- | _ =>
- default2
- tt
- end
- | None =>
- default1
- tt
- end
- | _ =>
- default1
- tt
- end in
- 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 :
+ ($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
+ =>
+ UnderLets.Base
(
#
(ident.fst)%expr @
- ($v)%expr,
- (-
- (#
+ ($v1)%expr,
+ #
(ident.snd)%expr @
- $v)%expr_pat)%expr)%expr_pat)
- else
- default
- | None =>
- default1
- tt
- end
- | _ =>
- default1
- tt
- end in
- 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) &&
- (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
+ ($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 @
- ($v)%expr,
- (-
- (#(ident.snd)%expr @
- $v)%expr_pat)%expr)%expr_pat)
- else default
- | None => default1 tt
- end
- | _ => default1 tt
- end in
- match x2 with
- | @expr.App _ _ _ s _ #(idc0)%expr_pat x3 =>
+ =>
+ 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 idc0 with
- | ident.Z_opp => Some tt
+ 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 _ =>
- match
- s as t2
- return
- (Compile.value' false t2 ->
- UnderLets.UnderLets base.type ident
- var
- (defaults.expr
+ 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
+ (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)))
- 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
+ (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
+ (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
+ | _ =>
+ 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
+ (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
+ | 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)))
- (fun
- a : option
+ 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
+ (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 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 t2) ->
+ (type.base t3) ->
defaults.expr
(type.base base.type.Z))
- =>
- match a with
- | Some x' =>
+ =>
+ 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_sub_get_borrow)%expr @
- x @ x1 @ x' v)%expr_pat
+ (#(ident.Z_add_get_carry)%expr @
+ x @ x1 @ x2)%expr_pat
(fun
- v0 : var
+ v2 : 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)
+ ($v2)%expr,
+ #(ident.snd)%expr @
+ ($v2)%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
+ 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)))
- (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 default
- | None => default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value'
- false s0 ->
- Compile.value'
- true d0 =>
- default
- end (Compile.reflect x3)
- | None => default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 =>
- default
- end (Compile.reflect x3)
- | None => default
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _
- _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s _ (_ @ _)%expr_pat _ | @expr.App _ _ _
- s _ (@expr.LetIn _ _ _ _ _ _ _) _ => default
- | _ => default1 tt
- end in
- match x1 with
+ 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 =>
+ 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 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
@@ -4434,83 +9270,387 @@ match idc in (ident t) return (Compile.value' true t) with
($v1)%expr,
(-
(#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- else default
- | None => default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 => default
- end (Compile.reflect x3)
- | None => default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value' false s0 ->
- Compile.value' true d0 => default
- end (Compile.reflect x3)
- | None => default
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => default
- | _ => default1 tt
- end
- | None => default0 tt
- end
- | @expr.App _ _ _ s _ #(idc)%expr_pat x3 =>
- match match idc with
- | ident.Z_opp => Some tt
- | _ => None
- end with
- | Some _ =>
- llet default1 := fun 'tt =>
- llet default1 := fun 'tt =>
- llet default1 := fun 'tt =>
- match x2 with
- | #(idc0)%expr_pat =>
- match
- match
- idc0
- with
- | @ident.Literal
- t1 v =>
- match
- t1 as t2
- return
- (base.base_interp
- t2 ->
+ 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
+ with
+ | base.type.unit =>
+ fun
_ : unit
- => None
- | base.type.Z =>
- fun
- v0 : Z =>
- Some v0
- | base.type.bool =>
- fun
+ => None
+ | base.type.Z =>
+ fun v2 : Z
+ =>
+ Some v2
+ | base.type.bool =>
+ fun
_ : bool
- => None
- | base.type.nat =>
- fun
+ => None
+ | base.type.nat =>
+ fun
_ : nat
- => None
- end v
- | _ =>
- None
- end
- with
- | Some
- args0 =>
- match
- s as t2
- return
+ => 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 t2 ->
+ false t5 ->
UnderLets.UnderLets
base.type
ident var
@@ -4518,21 +9658,21 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
(base.type.Z *
base.type.Z)%etype)))
- with
- | type.base
- t2 =>
+ with
+ | type.base
+ t5 =>
fun
- v :
+ v2 :
defaults.expr
(type.base
- t2) =>
+ t5) =>
base.try_make_transport_cps
(fun
- t1 : base.type
+ t6 : base.type
=>
defaults.expr
(type.base
- t1)) t2
+ t6)) t5
base.type.Z
(UnderLets.UnderLets
base.type
@@ -4542,33 +9682,33 @@ match idc in (ident t) return (Compile.value' true t) with
(base.type.Z *
base.type.Z)%etype)))
(fun
- a :
+ a2 :
option
(defaults.expr
(type.base
- t2) ->
+ t5) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a
+ a2
with
| Some
- x' =>
+ x'2 =>
if
- args0 <=?
- 0
+ args <? 0
then
UnderLets.UnderLet
(#
(ident.Z_sub_with_get_borrow)%expr @
- x @
- x' v @ x1 @
+ x @
##
- ((- args0)%Z)%expr)%expr_pat
+ ((- args)%Z)%expr @
+ x1 @
+ x'2 v2)%expr_pat
(fun
- v0 :
+ v3 :
var
(type.base
(base.type.Z *
@@ -4578,118 +9718,22 @@ match idc in (ident t) return (Compile.value' true t) with
(
#
(ident.fst)%expr @
- ($v0)%expr,
+ ($v3)%expr,
(-
(#
(ident.snd)%expr @
- $v0)%expr_pat)%expr)%expr_pat)
+ $v3)%expr_pat)%expr)%expr_pat)
else
- default
- | None =>
- default
- end)
- | (s0 ->
- d0)%ptype =>
- fun
- _ :
- Compile.value'
- false s0 ->
- Compile.value'
- true d0
- =>
- default
- end
- (Compile.reflect
- x3)
- | None =>
- default
- end
- | _ => default
- end in
- 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
- 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
+ if
+ args =? 0
+ then
+ UnderLets.UnderLet
(#
- (ident.Z_sub_with_get_borrow)%expr @
- x @
- x' v @ x2 @
- ##
- ((- args0)%Z)%expr)%expr_pat
+ (ident.Z_add_get_carry)%expr @
+ x @ x1 @
+ x2)%expr_pat
(fun
- v0 :
+ v3 :
var
(type.base
(base.type.Z *
@@ -4699,151 +9743,4822 @@ match idc in (ident t) return (Compile.value' true t) with
(
#
(ident.fst)%expr @
- ($v0)%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 @
- $v0)%expr_pat)%expr)%expr_pat)
- else default
- | None =>
- default
- end)
- | (s0 -> d0)%ptype =>
- fun
- _ : Compile.value'
- false s0 ->
- Compile.value'
- true d0 =>
- default
- end
- (Compile.reflect x3)
- | None => default1 tt
- end
- | _ => default1 tt
- end in
- match x2 with
- | @expr.App _ _ _ s0 _ #(idc0)%expr_pat x4 =>
+ ($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 idc0 with
- | ident.Z_opp => Some tt
+ 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 _ =>
- 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
+ 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)))
- (fun
- a : option
+ 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
+ (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 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 t2) ->
+ (type.base t3) ->
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)
+ =>
+ 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)
=>
- 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
+ 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)))
+ 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
+ (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 _ #(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
- a0 : option
- (defaults.expr
- (type.base
- t3) ->
- defaults.expr
- (type.base
- base.type.Z))
+ v1 : var
+ (type.base
+ (base.type.Z *
+ base.type.Z)%etype)
=>
- 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
+ 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,
- (-
- (#(ident.snd)%expr @
- $v1)%expr_pat)%expr)%expr_pat)
- | None => default
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
- false s1 ->
- Compile.value' true
- d1 => default
- end (Compile.reflect x4)
- | None => default
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- default
- end (Compile.reflect x3)
- | None => default
- end
- | @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _
- _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App
- _ _ _ s0 _ (_ @ _)%expr_pat _ | @expr.App _ _
- _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
- default
- | _ => default1 tt
- end in
- match x1 with
+ =>
+ 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
@@ -4914,39 +14629,2041 @@ match idc in (ident t) return (Compile.value' true t) with
($v1)%expr,
(-
(#(ident.snd)%expr @ $v1)%expr_pat)%expr)%expr_pat)
- | None => default
+ | 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 => default
+ Compile.value' true d1 =>
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @
+ x0 @ x1 @ x2)%expr_pat
end (Compile.reflect x4)
- | None => default
+ | 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 => default
+ Compile.value' true d1 =>
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
+ x2)%expr_pat
end (Compile.reflect x3)
- | None => default
+ | 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
+ (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 _ | @expr.App _ _ _ s0 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s0 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => default
- | _ => default1 tt
+ | 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
- | None => default
+ | ($_)%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
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => default
- | _ => default0 tt
end
| ident.Z_sub_get_borrow =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet _ := UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat in
UnderLets.UnderLet
(#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
(fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
@@ -4954,8 +16671,6 @@ match idc in (ident t) return (Compile.value' true t) with
(#(ident.fst)%expr @ ($v)%expr, #(ident.snd)%expr @ ($v)%expr)%expr_pat)
| ident.Z_sub_with_get_borrow =>
fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- llet _ := UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat in
UnderLets.UnderLet
(#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
(fun v : var (type.base (base.type.Z * base.type.Z)%etype) =>
@@ -4963,30 +16678,21 @@ match idc in (ident t) return (Compile.value' true t) with
(#(ident.fst)%expr @ ($v)%expr, #(ident.snd)%expr @ ($v)%expr)%expr_pat)
| ident.Z_zselect =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
| ident.Z_add_modulo =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
| ident.Z_rshi =>
fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
| ident.Z_cc_m =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat
| ident.Z_cast range =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_cast range)%expr @ x)%expr_pat in
- default
+ 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) =>
- llet default := UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat in
match x with
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x1) x0 =>
match
@@ -5049,103 +16755,88 @@ match idc in (ident t) return (Compile.value' true t) with
#(ident.Z_cast (snd range))%expr @
($(x'0 v0))%expr)%expr_pat;
UnderLets.Base (id (id fv)))%under_lets
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_cast2 range)%expr @ x)%expr_pat
end)
| (s1 -> d1)%ptype =>
fun
_ : Compile.value' false s1 ->
- Compile.value' true d1 => default
+ Compile.value' true d1 =>
+ UnderLets.Base
+ (#(ident.Z_cast2 range)%expr @ x)%expr_pat
end (Compile.reflect x0)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_cast2 range)%expr @ x)%expr_pat
end)
| (s1 -> d1)%ptype =>
fun _ : Compile.value' false s1 -> Compile.value' true d1 =>
- default
+ UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
end (Compile.reflect x1)
- | None => default
+ | 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 _ _ _ _ _ _ _) _) _ => default
+ 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 _ _ _ _ _ _ _) _ => default
- | _ => default
+ (@expr.LetIn _ _ _ _ _ _ _) _ =>
+ UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
+ | _ => UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
end
| ident.fancy_add log2wordmax imm =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
| ident.fancy_addc log2wordmax imm =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
| ident.fancy_sub log2wordmax imm =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
| ident.fancy_subb log2wordmax imm =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
| ident.fancy_mulll log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
| ident.fancy_mullh log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
| ident.fancy_mulhl log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
| ident.fancy_mulhh log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
| ident.fancy_rshi log2wordmax x =>
fun x0 : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
| ident.fancy_selc =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat
| ident.fancy_selm log2wordmax =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
| ident.fancy_sell =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat
| ident.fancy_addm =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat
end
: Compile.value' true t
diff --git a/src/Experiments/NewPipeline/fancy_rewrite_head.out b/src/Experiments/NewPipeline/fancy_rewrite_head.out
index 5d5edd8e9..b93038b7c 100644
--- a/src/Experiments/NewPipeline/fancy_rewrite_head.out
+++ b/src/Experiments/NewPipeline/fancy_rewrite_head.out
@@ -8,62 +8,43 @@ 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 => llet default := UnderLets.Base ##(v0)%expr in
- default
- | base.type.Z =>
- fun v0 : Z => llet default := UnderLets.Base ##(v0)%expr in
- default
- | base.type.bool =>
- fun v0 : bool => llet default := UnderLets.Base ##(v0)%expr in
- default
- | base.type.nat =>
- fun v0 : nat => llet default := UnderLets.Base ##(v0)%expr in
- default
+ | 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) =>
- llet default := UnderLets.Base (#(ident.Nat_succ)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Nat_succ)%expr @ x)%expr_pat
| ident.Nat_pred =>
fun x : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_pred)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Nat_pred)%expr @ x)%expr_pat
| ident.Nat_max =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat
| ident.Nat_mul =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat
| ident.Nat_add =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat
| ident.Nat_sub =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat in
- default
-| @ident.nil t => llet default := UnderLets.Base []%expr_pat in
- default
+ UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat
+| @ident.nil t => UnderLets.Base []%expr_pat
| @ident.cons t =>
fun (x : defaults.expr (type.base t))
(x0 : defaults.expr (type.base (base.type.list t))) =>
- llet default := UnderLets.Base (x :: x0)%expr_pat in
- default
+ UnderLets.Base (x :: x0)%expr_pat
| @ident.pair A B =>
fun (x : defaults.expr (type.base A)) (x0 : defaults.expr (type.base B))
- => llet default := UnderLets.Base (x, x0)%expr_pat in
- default
+ => UnderLets.Base (x, x0)%expr_pat
| @ident.fst A B =>
fun x : defaults.expr (type.base (A * B)%etype) =>
- llet default := UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
| @ident.snd A B =>
fun x : defaults.expr (type.base (A * B)%etype) =>
- llet default := UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
| @ident.prod_rect A B T =>
fun
(x : defaults.expr (type.base A) ->
@@ -71,11 +52,10 @@ 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)) =>
- llet default := 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 in
- default
+ 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
| @ident.bool_rect T =>
fun
(x
@@ -83,13 +63,12 @@ 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)) =>
- llet default := 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 in
- default
+ 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
| @ident.nat_rect P =>
fun
(x : defaults.expr (type.base base.type.unit) ->
@@ -100,15 +79,12 @@ 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)) =>
- llet default := 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 in
- default
+ 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
| @ident.nat_rect_arrow P Q =>
fun
(x : defaults.expr (type.base P) ->
@@ -123,22 +99,18 @@ 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)) =>
- llet default := 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 in
- default
+ 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
| @ident.list_rect A P =>
fun
(x : defaults.expr (type.base base.type.unit) ->
@@ -150,19 +122,13 @@ 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))) =>
- llet default := 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 in
- default
+ 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
| @ident.list_case A P =>
fun
(x : defaults.expr (type.base base.type.unit) ->
@@ -173,89 +139,70 @@ 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))) =>
- llet default := 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 in
- default
+ 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
| @ident.List_length T =>
fun x : defaults.expr (type.base (base.type.list T)) =>
- llet default := UnderLets.Base (#(ident.List_length)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_length)%expr @ x)%expr_pat
| ident.List_seq =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat
| @ident.List_firstn A =>
fun (x : defaults.expr (type.base base.type.nat))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- llet default := UnderLets.Base
- (#(ident.List_firstn)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_firstn)%expr @ x @ x0)%expr_pat
| @ident.List_skipn A =>
fun (x : defaults.expr (type.base base.type.nat))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- llet default := UnderLets.Base
- (#(ident.List_skipn)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_skipn)%expr @ x @ x0)%expr_pat
| @ident.List_repeat A =>
fun (x : defaults.expr (type.base A))
(x0 : defaults.expr (type.base base.type.nat)) =>
- llet default := UnderLets.Base
- (#(ident.List_repeat)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_repeat)%expr @ x @ x0)%expr_pat
| @ident.List_combine A B =>
fun (x : defaults.expr (type.base (base.type.list A)))
(x0 : defaults.expr (type.base (base.type.list B))) =>
- llet default := UnderLets.Base
- (#(ident.List_combine)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_combine)%expr @ x @ x0)%expr_pat
| @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))) =>
- llet default := UnderLets.Base
- (#(ident.List_map)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat in
- default
+ UnderLets.Base
+ (#(ident.List_map)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
| @ident.List_app A =>
fun x x0 : defaults.expr (type.base (base.type.list A)) =>
- llet default := UnderLets.Base (x ++ x0)%expr in
- default
+ UnderLets.Base (x ++ x0)%expr
| @ident.List_rev A =>
fun x : defaults.expr (type.base (base.type.list A)) =>
- llet default := UnderLets.Base (#(ident.List_rev)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.List_rev)%expr @ x)%expr_pat
| @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))) =>
- llet default := UnderLets.Base
- (#(ident.List_flat_map)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat in
- default
+ UnderLets.Base
+ (#(ident.List_flat_map)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
| @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))) =>
- llet default := UnderLets.Base
- (#(ident.List_partition)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat in
- default
+ UnderLets.Base
+ (#(ident.List_partition)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
| @ident.List_fold_right A B =>
fun
(x : defaults.expr (type.base B) ->
@@ -264,36 +211,30 @@ 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))) =>
- llet default := 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 in
- default
+ 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
| @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))) =>
- llet default := UnderLets.Base
- (#(ident.List_update_nth)%expr @ x @
- (λ x2 : var (type.base T),
- UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat in
- default
+ UnderLets.Base
+ (#(ident.List_update_nth)%expr @ x @
+ (λ x2 : var (type.base T),
+ UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat
| @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)) =>
- llet default := UnderLets.Base
- (#(ident.List_nth_default)%expr @ x @ x0 @ x1)%expr_pat in
- default
+ 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) =>
- llet default := UnderLets.Base (x + x0)%expr in
- default
+ UnderLets.Base (x + x0)%expr
| ident.Z_mul =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x * x0)%expr in
match x with
| #(idc)%expr_pat =>
match
@@ -319,68 +260,1284 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some _ =>
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- match x1 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal
- t2 v =>
+ 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
- t2 as t3
+ 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
- (base.base_interp
- t3 ->
- option Z)
+ (Compile.value'
+ false t5 ->
+ UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
+ base.type.Z)))
with
- | base.type.unit =>
+ | type.base
+ t5 =>
fun
- _ : unit
- => None
- | base.type.Z =>
+ 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
- v0 : Z =>
- Some v0
- | base.type.bool =>
+ 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
- _ : bool
- => None
- | base.type.nat =>
+ _ :
+ 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
- _ : nat
- => None
- end v
- | _ => None
+ 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
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value'
- false t2 ->
- UnderLets.UnderLets
+ 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
- t2 =>
- fun
- v :
+ with
+ | type.base
+ t5 =>
+ fun
+ v1 :
defaults.expr
(type.base
- t2) =>
- base.try_make_transport_cps
+ t5) =>
+ base.try_make_transport_cps
(fun
- t3 : base.type
+ t6 : base.type
=>
defaults.expr
(type.base
- t3)) t2
+ t6)) t5
base.type.Z
(UnderLets.UnderLets
base.type
@@ -389,20 +1546,96 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a :
+ a1 :
option
(defaults.expr
(type.base
- t2) ->
+ t5) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a
+ a1
with
| Some
- x' =>
+ 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
@@ -429,157 +1662,276 @@ match idc in (ident t) return (Compile.value' true t) with
args1))%expr @
(##
(x3)%expr,
- x' v))%expr_pat
+ x'2 v2))%expr_pat
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
else
- default
+ UnderLets.Base
+ (x * x0)%expr
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
- | (s1 -> d1)%ptype =>
- fun
+ | (s1 ->
+ d1)%ptype =>
+ fun
_ :
Compile.value'
false s1 ->
Compile.value'
true d1
- => default
- end
- (Compile.reflect
- x2)
- | None => default
- end
- | _ => default
- end in
- 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
+ =>
+ 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' v))%expr_pat
- | None =>
- match
- s as t3
- return
- (Compile.value'
- false t3 ->
- UnderLets.UnderLets
+ (##
+ (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
- t3 =>
- fun
- v0 :
+ with
+ | type.base
+ t6 =>
+ fun
+ v2 :
defaults.expr
(type.base
- t3) =>
- base.try_make_transport_cps
- (fun
- t4 : base.type
+ t6) =>
+ base.try_make_transport_cps
+ (fun
+ t7 : base.type
=>
defaults.expr
(type.base
- t4)) t3
- base.type.Z
- (UnderLets.UnderLets
+ t7)) t6
+ base.type.Z
+ (UnderLets.UnderLets
base.type
ident var
(defaults.expr
(type.base
base.type.Z)))
- (fun
- a0 :
+ (fun
+ a2 :
option
(defaults.expr
(type.base
- t3) ->
+ t6) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a0
+ a2
with
| Some
- x'0 =>
+ x'2 =>
if
args1 =?
2
@@ -606,71 +1958,82 @@ match idc in (ident t) return (Compile.value' true t) with
args1))%expr @
(##
(x3)%expr,
- x'0 v0))%expr_pat
+ x'2 v2))%expr_pat
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
else
- default
+ UnderLets.Base
+ (x * x0)%expr
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
- | (s1 -> d1)%ptype =>
- fun
- _ :
+ | (s1 ->
+ d1)%ptype =>
+ fun
+ _ :
Compile.value'
false s1 ->
Compile.value'
true d1
- => default
- 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
+ =>
+ 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 t3 =>
- fun
- v0 : defaults.expr
+ with
+ | type.base
+ t6 =>
+ fun
+ v2 :
+ 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
+ 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
+ (fun
+ a2 :
+ option
(defaults.expr
(type.base
- t3) ->
+ t6) ->
defaults.expr
(type.base
base.type.Z))
- =>
- match a0 with
- | Some x'0 =>
+ =>
+ match
+ a2
+ with
+ | Some
+ x'2 =>
if
args1 =?
2
@@ -697,104 +2060,1586 @@ match idc in (ident t) return (Compile.value' true t) with
args1))%expr @
(##
(x3)%expr,
- x'0 v0))%expr_pat
+ x'2 v2))%expr_pat
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
else
- default
- | None =>
- default
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value'
+ 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
- => default
- end
- (Compile.reflect
- x1)
- | None => default
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false
- s1 ->
- Compile.value' true d1
- => default
- end (Compile.reflect x1)
- | None => default0 tt
- end
- | _ => default0 tt
- end in
- 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 => default
+ =>
+ 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
- else default
- | None => default
+ | None => UnderLets.Base (x * x0)%expr
end)
| (s1 -> d1)%ptype =>
fun
_ : Compile.value' false s1 ->
- Compile.value' true d1 => default
+ Compile.value' true d1 =>
+ UnderLets.Base (x * x0)%expr
end (Compile.reflect x2)
- | None => default0 tt
+ | 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
+ 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' =>
+ 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
+ (#(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
- | _ => default0 tt
end
| None =>
match
@@ -902,30 +3747,37 @@ match idc in (ident t) return (Compile.value' true t) with
args1))%expr @
(##(x3)%expr,
x'0 v0))%expr_pat
- | None => default
+ | None =>
+ UnderLets.Base
+ (x * x0)%expr
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s1 -> d1)%ptype =>
fun
_ : Compile.value' false
s1 ->
Compile.value' true
- d1 => default
+ d1 =>
+ UnderLets.Base
+ (x * x0)%expr
end (Compile.reflect x2)
end
- | None => default
+ | None => UnderLets.Base (x * x0)%expr
end)
| (s1 -> d1)%ptype =>
fun
_ : Compile.value' false s1 ->
- Compile.value' true d1 => default
+ Compile.value' true d1 =>
+ UnderLets.Base (x * x0)%expr
end (Compile.reflect x2)
- | None => default
+ | None => UnderLets.Base (x * x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x * x0)%expr
end
- | None => default
+ | None => UnderLets.Base (x * x0)%expr
end
end
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ |
@@ -934,13 +3786,14 @@ match idc in (ident t) return (Compile.value' true t) with
_ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ |
@expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ =>
- default
+ UnderLets.Base (x * x0)%expr
| @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _
($_)%expr _ | @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => default
- | _ => default
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ =>
+ UnderLets.Base (x * x0)%expr
+ | _ => UnderLets.Base (x * x0)%expr
end
- | None => default
+ | None => UnderLets.Base (x * x0)%expr
end
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x2) x1 =>
match match idc with
@@ -948,34 +3801,2252 @@ match idc in (ident t) return (Compile.value' true t) with
| _ => None
end with
| Some _ =>
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- llet default0 :=
- fun 'tt =>
- match x1 with
- | #(idc0)%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 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
- idc0
+ invert_high
+ (2 *
+ Z.log2_up
+ args0)
+ args1
with
- | @ident.Literal
- t1 v =>
+ | Some
+ y =>
+ UnderLets.Base
+ (#
+ (ident.fancy_mullh
+ (2 *
+ Z.log2_up
+ args0))%expr @
+ (
+ x'1 v1,
+ ##
+ (y)%expr))%expr_pat
+ | None =>
match
- t1 as t2
+ 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
+ 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
- t2 ->
+ t7 ->
option Z)
with
| base.type.unit =>
@@ -984,8 +6055,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v2 : Z =>
+ Some v2
| base.type.bool =>
fun
_ : bool
@@ -994,48 +6065,150 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
- | _ =>
- None
- end
+ 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
- | Some
- args0 =>
+ | 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
- x0
+ a1
with
- | (@expr.App
- _ _ _ s2
- _ #(idc1)
- x4 @ x3)%expr_pat =>
- match
+ | Some
+ x'1 =>
match
- idc1
+ s1 as t7
+ return
+ (Compile.value'
+ false t7 ->
+ UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
+ base.type.Z)))
with
- | ident.Z_shiftr =>
- Some tt
- | _ =>
- None
- end
+ | 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
| #
- (idc2)%expr_pat =>
+ (idc4)%expr_pat =>
match
match
- idc2
+ idc4
with
| @ident.Literal
- t3 v =>
+ t9 v3 =>
match
- t3 as t4
+ t9 as t10
return
(base.base_interp
- t4 ->
+ t10 ->
option Z)
with
| base.type.unit =>
@@ -1044,8 +6217,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v4 : Z =>
+ Some v4
| base.type.bool =>
fun
_ : bool
@@ -1054,18 +6227,18 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
+ end v3
| _ =>
None
end
with
| Some
- args2 =>
+ args4 =>
match
- s0 as t3
+ s as t9
return
(Compile.value'
- false t3 ->
+ false t9 ->
UnderLets.UnderLets
base.type
ident var
@@ -1074,19 +6247,19 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)))
with
| type.base
- t3 =>
+ t9 =>
fun
- v :
+ v3 :
defaults.expr
(type.base
- t3) =>
+ t9) =>
base.try_make_transport_cps
(fun
- t4 : base.type
+ t10 : base.type
=>
defaults.expr
(type.base
- t4)) t3
+ t10)) t9
base.type.Z
(UnderLets.UnderLets
base.type
@@ -1095,25 +6268,25 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a :
+ a3 :
option
(defaults.expr
(type.base
- t3) ->
+ t9) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a
+ a3
with
| Some
- x' =>
+ x'3 =>
match
- s2 as t4
+ s2 as t10
return
(Compile.value'
- false t4 ->
+ false t10 ->
UnderLets.UnderLets
base.type
ident var
@@ -1122,19 +6295,19 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)))
with
| type.base
- t4 =>
+ t10 =>
fun
- v0 :
+ v4 :
defaults.expr
(type.base
- t4) =>
+ t10) =>
base.try_make_transport_cps
(fun
- t5 : base.type
+ t11 : base.type
=>
defaults.expr
(type.base
- t5)) t4
+ t11)) t10
base.type.Z
(UnderLets.UnderLets
base.type
@@ -1143,40 +6316,173 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a0 :
+ a4 :
option
(defaults.expr
(type.base
- t4) ->
+ t10) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a0
+ a4
with
| Some
- x'0 =>
+ x'4 =>
if
- args0 =?
+ (args0 =?
2
^
(2 *
- args2 / 2) -
- 1
+ Z.log2_up
+ args0 / 2) -
+ 1) &&
+ (args4 =?
+ 2
+ ^
+ (2 *
+ Z.log2_up
+ args0 / 2) -
+ 1)
then
UnderLets.Base
(#
- (ident.fancy_mullh
+ (ident.fancy_mulll
(2 *
- args2))%expr @
+ Z.log2_up
+ args0))%expr @
(
- x' v,
- x'0 v0))%expr_pat
+ 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
- default
+ UnderLets.Base
+ (x * x0)%expr
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 ->
d3)%ptype =>
@@ -1187,12 +6493,14 @@ match idc in (ident t) return (Compile.value' true t) with
Compile.value'
true d3
=>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect
x4)
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 ->
d3)%ptype =>
@@ -1203,64 +6511,108 @@ match idc in (ident t) return (Compile.value' true t) with
Compile.value'
true d3
=>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect
x2)
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
+ end)
+ | (s3 ->
+ d3)%ptype =>
+ fun
+ _ :
+ Compile.value'
+ false s3 ->
+ Compile.value'
+ true d3
+ =>
+ UnderLets.Base
+ (x * x0)%expr
end
- | _ =>
- default
+ (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 =>
- default
+ UnderLets.Base
+ (x * x0)%expr
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 =>
- default
| _ =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
- | None =>
- default
+ | 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
- | _ => default
- end in
- match x2 with
- | #(idc0)%expr_pat =>
+ (Compile.reflect
+ x2)
+ | None =>
match
+ x3
+ with
+ | #
+ (idc4)%expr_pat =>
+ match
match
- idc0
+ idc4
with
| @ident.Literal
- t1 v =>
+ t7 v1 =>
match
- t1 as t2
+ t7 as t8
return
(base.base_interp
- t2 ->
+ t8 ->
option Z)
with
| base.type.unit =>
@@ -1269,8 +6621,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v2 : Z =>
+ Some v2
| base.type.bool =>
fun
_ : bool
@@ -1279,48 +6631,196 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
+ end v1
| _ =>
None
end
- with
- | Some
- args0 =>
- match
- x0
with
- | (@expr.App
- _ _ _ s2
- _ #(idc1)
- x4 @ x3)%expr_pat =>
- match
+ | Some
+ args3 =>
match
- idc1
+ s as t7
+ return
+ (Compile.value'
+ false t7 ->
+ UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
+ base.type.Z)))
with
- | ident.Z_shiftr =>
- Some tt
- | _ =>
- None
- end
+ | 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
- x3
+ s2 as t8
+ return
+ (Compile.value'
+ false t8 ->
+ UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
+ base.type.Z)))
with
- | #
- (idc2)%expr_pat =>
+ | 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
- idc2
+ idc3
with
| @ident.Literal
- t3 v =>
+ t7 v2 =>
match
- t3 as t4
+ t7 as t8
return
(base.base_interp
- t4 ->
+ t8 ->
option Z)
with
| base.type.unit =>
@@ -1329,8 +6829,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v3 : Z =>
+ Some v3
| base.type.bool =>
fun
_ : bool
@@ -1339,18 +6839,18 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
+ end v2
| _ =>
None
end
- with
- | Some
- args2 =>
+ with
+ | Some
+ args3 =>
match
- s as t3
+ s as t7
return
(Compile.value'
- false t3 ->
+ false t7 ->
UnderLets.UnderLets
base.type
ident var
@@ -1359,19 +6859,19 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)))
with
| type.base
- t3 =>
+ t7 =>
fun
- v :
+ v2 :
defaults.expr
(type.base
- t3) =>
+ t7) =>
base.try_make_transport_cps
(fun
- t4 : base.type
+ t8 : base.type
=>
defaults.expr
(type.base
- t4)) t3
+ t8)) t7
base.type.Z
(UnderLets.UnderLets
base.type
@@ -1380,25 +6880,25 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a :
+ a1 :
option
(defaults.expr
(type.base
- t3) ->
+ t7) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a
+ a1
with
| Some
- x' =>
+ x'1 =>
match
- s2 as t4
+ s2 as t8
return
(Compile.value'
- false t4 ->
+ false t8 ->
UnderLets.UnderLets
base.type
ident var
@@ -1407,19 +6907,19 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)))
with
| type.base
- t4 =>
+ t8 =>
fun
- v0 :
+ v3 :
defaults.expr
(type.base
- t4) =>
+ t8) =>
base.try_make_transport_cps
(fun
- t5 : base.type
+ t9 : base.type
=>
defaults.expr
(type.base
- t5)) t4
+ t9)) t8
base.type.Z
(UnderLets.UnderLets
base.type
@@ -1428,40 +6928,51 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a0 :
+ a2 :
option
(defaults.expr
(type.base
- t4) ->
+ t8) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a0
+ a2
with
| Some
- x'0 =>
+ x'2 =>
if
- args0 =?
+ (args0 =?
2
^
(2 *
- args2 / 2) -
- 1
+ Z.log2_up
+ args0 / 2) -
+ 1) &&
+ (args3 =?
+ 2
+ ^
+ (2 *
+ Z.log2_up
+ args0 / 2) -
+ 1)
then
UnderLets.Base
(#
- (ident.fancy_mullh
+ (ident.fancy_mulll
(2 *
- args2))%expr @
+ Z.log2_up
+ args0))%expr @
(
- x' v,
- x'0 v0))%expr_pat
+ x'1 v2,
+ x'2 v3))%expr_pat
else
- default
+ UnderLets.Base
+ (x * x0)%expr
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 ->
d3)%ptype =>
@@ -1472,12 +6983,14 @@ match idc in (ident t) return (Compile.value' true t) with
Compile.value'
true d3
=>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect
x4)
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 ->
d3)%ptype =>
@@ -1488,67 +7001,34 @@ match idc in (ident t) return (Compile.value' true t) with
Compile.value'
true d3
=>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect
x1)
- | None =>
- default
- end
- | _ =>
- default
- end
- | None =>
- default
- 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 =>
- default
- | _ =>
- default
- end
| None =>
- default0
- tt
+ UnderLets.Base
+ (x * x0)%expr
end
| _ =>
- default0
- tt
- end in
- match x1 with
- | #(idc0)%expr_pat =>
+ UnderLets.Base
+ (x * x0)%expr
+ end
+ | @expr.LetIn _ _
+ _ _ _ _ _ =>
+ match x3 with
+ | #(idc3)%expr_pat =>
match
match
- idc0
+ idc3
with
| @ident.Literal
- t1 v =>
+ t6 v1 =>
match
- t1 as t2
+ t6 as t7
return
(base.base_interp
- t2 ->
+ t7 ->
option Z)
with
| base.type.unit =>
@@ -1557,8 +7037,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v2 : Z =>
+ Some v2
| base.type.bool =>
fun
_ : bool
@@ -1567,48 +7047,553 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
+ end v1
| _ =>
None
end
with
| Some
- args0 =>
+ args3 =>
match
- x0
+ s as t6
+ return
+ (Compile.value'
+ false t6 ->
+ UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
+ base.type.Z)))
with
- | (@expr.App
- _ _ _ s2
- _ #(idc1)
- x4 @ x3)%expr_pat =>
+ | 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
- idc1
+ t6 as t7
+ return
+ (base.base_interp
+ t7 ->
+ option Z)
with
- | ident.Z_land =>
- Some tt
+ | 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
- x3
+ s2 as t7
+ return
+ (Compile.value'
+ false t7 ->
+ UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
+ base.type.Z)))
with
- | #
- (idc2)%expr_pat =>
+ | 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
- idc2
+ idc4
with
| @ident.Literal
- t3 v =>
+ t7 v1 =>
match
- t3 as t4
+ t7 as t8
return
(base.base_interp
- t4 ->
+ t8 ->
option Z)
with
| base.type.unit =>
@@ -1617,8 +7602,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v2 : Z =>
+ Some v2
| base.type.bool =>
fun
_ : bool
@@ -1627,18 +7612,18 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
+ end v1
| _ =>
None
end
with
| Some
- args2 =>
+ args3 =>
match
- s0 as t3
+ s0 as t7
return
(Compile.value'
- false t3 ->
+ false t7 ->
UnderLets.UnderLets
base.type
ident var
@@ -1647,19 +7632,19 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)))
with
| type.base
- t3 =>
+ t7 =>
fun
- v :
+ v1 :
defaults.expr
(type.base
- t3) =>
+ t7) =>
base.try_make_transport_cps
(fun
- t4 : base.type
+ t8 : base.type
=>
defaults.expr
(type.base
- t4)) t3
+ t8)) t7
base.type.Z
(UnderLets.UnderLets
base.type
@@ -1668,25 +7653,25 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a :
+ a1 :
option
(defaults.expr
(type.base
- t3) ->
+ t7) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a
+ a1
with
| Some
- x' =>
+ x'1 =>
match
- s2 as t4
+ s2 as t8
return
(Compile.value'
- false t4 ->
+ false t8 ->
UnderLets.UnderLets
base.type
ident var
@@ -1695,19 +7680,19 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)))
with
| type.base
- t4 =>
+ t8 =>
fun
- v0 :
+ v2 :
defaults.expr
(type.base
- t4) =>
+ t8) =>
base.try_make_transport_cps
(fun
- t5 : base.type
+ t9 : base.type
=>
defaults.expr
(type.base
- t5)) t4
+ t9)) t8
base.type.Z
(UnderLets.UnderLets
base.type
@@ -1716,34 +7701,34 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a0 :
+ a2 :
option
(defaults.expr
(type.base
- t4) ->
+ t8) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a0
+ a2
with
| Some
- x'0 =>
+ x'2 =>
if
- (args0 =?
+ (args3 =?
2
^
(2 *
Z.log2_up
- args0 / 2) -
+ args3 / 2) -
1) &&
(args2 =?
2
^
(2 *
Z.log2_up
- args0 / 2) -
+ args3 / 2) -
1)
then
UnderLets.Base
@@ -1751,14 +7736,16 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.fancy_mulll
(2 *
Z.log2_up
- args0))%expr @
+ args3))%expr @
(
- x' v,
- x'0 v0))%expr_pat
+ x'1 v1,
+ x'2 v2))%expr_pat
else
- default
+ UnderLets.Base
+ (x * x0)%expr
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 ->
d3)%ptype =>
@@ -1769,12 +7756,14 @@ match idc in (ident t) return (Compile.value' true t) with
Compile.value'
true d3
=>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect
x4)
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 ->
d3)%ptype =>
@@ -1785,67 +7774,179 @@ match idc in (ident t) return (Compile.value' true t) with
Compile.value'
true d3
=>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect
x2)
| None =>
- default
- end
- | _ =>
- default
- end
- | None =>
- default
- 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 =>
- default
- | _ =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
- | None =>
- default0
- tt
+ | _ =>
+ UnderLets.Base
+ (x * x0)%expr
end
- | _ =>
- default0
- tt
- end in
- match x2 with
- | #(idc0)%expr_pat =>
+ | 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
- idc0
+ idc3
with
| @ident.Literal
- t1 v =>
+ t7 v2 =>
match
- t1 as t2
+ t7 as t8
return
(base.base_interp
- t2 ->
+ t8 ->
option Z)
with
| base.type.unit =>
@@ -1854,8 +7955,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v3 : Z =>
+ Some v3
| base.type.bool =>
fun
_ : bool
@@ -1864,48 +7965,337 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
+ end v2
| _ =>
None
end
with
| Some
- args0 =>
+ args3 =>
match
- x0
+ s0 as t7
+ return
+ (Compile.value'
+ false t7 ->
+ UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
+ base.type.Z)))
with
- | (@expr.App
- _ _ _ s2
- _ #(idc1)
- x4 @ x3)%expr_pat =>
- match
+ | 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
- idc1
- with
- | ident.Z_land =>
- Some tt
- | _ =>
- None
- end
+ a1
with
| Some
- _ =>
+ x'1 =>
match
- x3
+ s2 as t8
+ return
+ (Compile.value'
+ false t8 ->
+ UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
+ base.type.Z)))
with
- | #
- (idc2)%expr_pat =>
+ | 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
- idc2
+ idc3
with
| @ident.Literal
- t3 v =>
+ t6 v1 =>
match
- t3 as t4
+ t6 as t7
return
(base.base_interp
- t4 ->
+ t7 ->
option Z)
with
| base.type.unit =>
@@ -1914,8 +8304,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v2 : Z =>
+ Some v2
| base.type.bool =>
fun
_ : bool
@@ -1924,18 +8314,18 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
+ end v1
| _ =>
None
end
- with
- | Some
- args2 =>
+ with
+ | Some
+ args3 =>
match
- s as t3
+ s0 as t6
return
(Compile.value'
- false t3 ->
+ false t6 ->
UnderLets.UnderLets
base.type
ident var
@@ -1944,19 +8334,19 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)))
with
| type.base
- t3 =>
+ t6 =>
fun
- v :
+ v1 :
defaults.expr
(type.base
- t3) =>
+ t6) =>
base.try_make_transport_cps
(fun
- t4 : base.type
+ t7 : base.type
=>
defaults.expr
(type.base
- t4)) t3
+ t7)) t6
base.type.Z
(UnderLets.UnderLets
base.type
@@ -1965,25 +8355,25 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a :
+ a1 :
option
(defaults.expr
(type.base
- t3) ->
+ t6) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a
+ a1
with
| Some
- x' =>
+ x'1 =>
match
- s2 as t4
+ s2 as t7
return
(Compile.value'
- false t4 ->
+ false t7 ->
UnderLets.UnderLets
base.type
ident var
@@ -1992,19 +8382,19 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)))
with
| type.base
- t4 =>
+ t7 =>
fun
- v0 :
+ v2 :
defaults.expr
(type.base
- t4) =>
+ t7) =>
base.try_make_transport_cps
(fun
- t5 : base.type
+ t8 : base.type
=>
defaults.expr
(type.base
- t5)) t4
+ t8)) t7
base.type.Z
(UnderLets.UnderLets
base.type
@@ -2013,34 +8403,34 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a0 :
+ a2 :
option
(defaults.expr
(type.base
- t4) ->
+ t7) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a0
+ a2
with
| Some
- x'0 =>
+ x'2 =>
if
- (args0 =?
+ (args3 =?
2
^
(2 *
Z.log2_up
- args0 / 2) -
+ args3 / 2) -
1) &&
(args2 =?
2
^
(2 *
Z.log2_up
- args0 / 2) -
+ args3 / 2) -
1)
then
UnderLets.Base
@@ -2048,14 +8438,16 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.fancy_mulll
(2 *
Z.log2_up
- args0))%expr @
+ args3))%expr @
(
- x' v,
- x'0 v0))%expr_pat
+ x'1 v1,
+ x'2 v2))%expr_pat
else
- default
+ UnderLets.Base
+ (x * x0)%expr
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 ->
d3)%ptype =>
@@ -2066,12 +8458,14 @@ match idc in (ident t) return (Compile.value' true t) with
Compile.value'
true d3
=>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect
x4)
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 ->
d3)%ptype =>
@@ -2082,67 +8476,175 @@ match idc in (ident t) return (Compile.value' true t) with
Compile.value'
true d3
=>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect
- x1)
- | None =>
- default
- end
- | _ =>
- default
- end
- | None =>
- default
- 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 =>
- default
- | _ =>
- default
- end
+ x2)
| None =>
- default0
- tt
+ UnderLets.Base
+ (x * x0)%expr
end
| _ =>
- default0
- tt
- end in
+ 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
+ (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
- | #(idc0)%expr_pat =>
+ | #(idc3)%expr_pat =>
match
match
- idc0
+ idc3
with
| @ident.Literal
- t1 v =>
+ t6 v1 =>
match
- t1 as t2
+ t6 as t7
return
(base.base_interp
- t2 ->
+ t7 ->
option Z)
with
| base.type.unit =>
@@ -2151,8 +8653,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v2 : Z =>
+ Some v2
| base.type.bool =>
fun
_ : bool
@@ -2161,51 +8663,336 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
+ end v1
| _ =>
None
end
with
| Some
- args0 =>
+ args3 =>
match
- x0
+ s0 as t6
+ return
+ (Compile.value'
+ false t6 ->
+ UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
+ base.type.Z)))
with
- | @expr.App
- _ _ _ s1
- _
- (#
- (idc1) @
- x4)%expr_pat
- x3 =>
- match
+ | 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
- idc1
- with
- | ident.Z_land =>
- Some tt
- | _ =>
- None
- end
+ a1
with
| Some
- _ =>
+ x'1 =>
match
- x4
+ s2 as t7
+ return
+ (Compile.value'
+ false t7 ->
+ UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
+ base.type.Z)))
with
- | #
- (idc2)%expr_pat =>
+ | 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
- idc2
+ idc3
with
| @ident.Literal
- t3 v =>
+ t6 v1 =>
match
- t3 as t4
+ t6 as t7
return
(base.base_interp
- t4 ->
+ t7 ->
option Z)
with
| base.type.unit =>
@@ -2214,8 +9001,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v2 : Z =>
+ Some v2
| base.type.bool =>
fun
_ : bool
@@ -2224,18 +9011,18 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
+ end v1
| _ =>
None
end
- with
- | Some
- args2 =>
+ with
+ | Some
+ args3 =>
match
- s0 as t3
+ s0 as t6
return
(Compile.value'
- false t3 ->
+ false t6 ->
UnderLets.UnderLets
base.type
ident var
@@ -2244,19 +9031,19 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)))
with
| type.base
- t3 =>
+ t6 =>
fun
- v :
+ v1 :
defaults.expr
(type.base
- t3) =>
+ t6) =>
base.try_make_transport_cps
(fun
- t4 : base.type
+ t7 : base.type
=>
defaults.expr
(type.base
- t4)) t3
+ t7)) t6
base.type.Z
(UnderLets.UnderLets
base.type
@@ -2265,25 +9052,25 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a :
+ a1 :
option
(defaults.expr
(type.base
- t3) ->
+ t6) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a
+ a1
with
| Some
- x' =>
+ x'1 =>
match
- s1 as t4
+ s2 as t7
return
(Compile.value'
- false t4 ->
+ false t7 ->
UnderLets.UnderLets
base.type
ident var
@@ -2292,19 +9079,19 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)))
with
| type.base
- t4 =>
+ t7 =>
fun
- v0 :
+ v2 :
defaults.expr
(type.base
- t4) =>
+ t7) =>
base.try_make_transport_cps
(fun
- t5 : base.type
+ t8 : base.type
=>
defaults.expr
(type.base
- t5)) t4
+ t8)) t7
base.type.Z
(UnderLets.UnderLets
base.type
@@ -2313,49 +9100,42 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a0 :
+ a2 :
option
(defaults.expr
(type.base
- t4) ->
+ t7) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a0
+ a2
with
| Some
- x'0 =>
+ x'2 =>
if
- (args0 =?
+ args3 =?
2
^
(2 *
- Z.log2_up
- args0 / 2) -
- 1) &&
- (args2 =?
- 2
- ^
- (2 *
- Z.log2_up
- args0 / 2) -
- 1)
+ args2 / 2) -
+ 1
then
UnderLets.Base
(#
- (ident.fancy_mulll
+ (ident.fancy_mullh
(2 *
- Z.log2_up
- args0))%expr @
+ args2))%expr @
(
- x' v,
- x'0 v0))%expr_pat
+ x'1 v1,
+ x'2 v2))%expr_pat
else
- default
+ UnderLets.Base
+ (x * x0)%expr
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 ->
d3)%ptype =>
@@ -2366,12 +9146,14 @@ match idc in (ident t) return (Compile.value' true t) with
Compile.value'
true d3
=>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect
- x3)
+ x4)
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 ->
d3)%ptype =>
@@ -2382,92 +9164,480 @@ match idc in (ident t) return (Compile.value' true t) with
Compile.value'
true d3
=>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect
x2)
- | None =>
- default
- end
- | _ =>
- default
- end
- | None =>
- default
- 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
- _ _ _ _ _
- _ _) _ =>
- default
- | _ =>
- default
- end
| None =>
- default0
- tt
+ UnderLets.Base
+ (x * x0)%expr
end
| _ =>
- default0
- tt
- end in
- match x2 with
- | #(idc0)%expr_pat =>
+ 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)
+ =>
+ 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
- idc0
+ match
+ idc4
with
| @ident.Literal
- t1 v =>
+ t7 v1 =>
match
- t1 as t2
+ t7 as t8
return
(base.base_interp
- t2 ->
+ t8 ->
option Z)
with
| base.type.unit =>
@@ -2476,8 +9646,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v2 : Z =>
+ Some v2
| base.type.bool =>
fun
_ : bool
@@ -2486,51 +9656,1468 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
+ end v1
| _ =>
None
end
- with
- | Some
- args0 =>
+ with
+ | Some
+ args3 =>
match
- x0
+ s0 as t7
+ return
+ (Compile.value'
+ false t7 ->
+ UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
+ base.type.Z)))
with
- | @expr.App
- _ _ _ s1
- _
- (#
- (idc1) @
- x4)%expr_pat
- x3 =>
+ | 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
- idc1
+ s2 as t8
+ return
+ (Compile.value'
+ false t8 ->
+ UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
+ base.type.Z)))
with
- | ident.Z_land =>
- Some tt
- | _ =>
- None
- end
+ | 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
- x4
+ a0
with
- | #
- (idc2)%expr_pat =>
- match
+ | 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
+ (#(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
+ (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
- idc2
+ idc3
with
| @ident.Literal
- t3 v =>
+ t7 v2 =>
match
- t3 as t4
+ t7 as t8
return
(base.base_interp
- t4 ->
+ t8 ->
option Z)
with
| base.type.unit =>
@@ -2539,8 +11126,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v3 : Z =>
+ Some v3
| base.type.bool =>
fun
_ : bool
@@ -2549,18 +11136,18 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
+ end v2
| _ =>
None
end
- with
- | Some
- args2 =>
+ with
+ | Some
+ args3 =>
match
- s as t3
+ s0 as t7
return
(Compile.value'
- false t3 ->
+ false t7 ->
UnderLets.UnderLets
base.type
ident var
@@ -2569,19 +11156,19 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)))
with
| type.base
- t3 =>
+ t7 =>
fun
- v :
+ v2 :
defaults.expr
(type.base
- t3) =>
+ t7) =>
base.try_make_transport_cps
(fun
- t4 : base.type
+ t8 : base.type
=>
defaults.expr
(type.base
- t4)) t3
+ t8)) t7
base.type.Z
(UnderLets.UnderLets
base.type
@@ -2590,25 +11177,25 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a :
+ a1 :
option
(defaults.expr
(type.base
- t3) ->
+ t7) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a
+ a1
with
| Some
- x' =>
+ x'1 =>
match
- s1 as t4
+ s2 as t8
return
(Compile.value'
- false t4 ->
+ false t8 ->
UnderLets.UnderLets
base.type
ident var
@@ -2617,19 +11204,19 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)))
with
| type.base
- t4 =>
+ t8 =>
fun
- v0 :
+ v3 :
defaults.expr
(type.base
- t4) =>
+ t8) =>
base.try_make_transport_cps
(fun
- t5 : base.type
+ t9 : base.type
=>
defaults.expr
(type.base
- t5)) t4
+ t9)) t8
base.type.Z
(UnderLets.UnderLets
base.type
@@ -2638,20 +11225,20 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a0 :
+ a2 :
option
(defaults.expr
(type.base
- t4) ->
+ t8) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a0
+ a2
with
| Some
- x'0 =>
+ x'2 =>
if
(args0 =?
2
@@ -2660,7 +11247,7 @@ match idc in (ident t) return (Compile.value' true t) with
Z.log2_up
args0 / 2) -
1) &&
- (args2 =?
+ (args3 =?
2
^
(2 *
@@ -2675,12 +11262,14 @@ match idc in (ident t) return (Compile.value' true t) with
Z.log2_up
args0))%expr @
(
- x' v,
- x'0 v0))%expr_pat
+ x'1 v2,
+ x'2 v3))%expr_pat
else
- default
+ UnderLets.Base
+ (x * x0)%expr
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 ->
d3)%ptype =>
@@ -2691,12 +11280,14 @@ match idc in (ident t) return (Compile.value' true t) with
Compile.value'
true d3
=>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect
- x3)
+ x4)
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 ->
d3)%ptype =>
@@ -2707,92 +11298,1242 @@ match idc in (ident t) return (Compile.value' true t) with
Compile.value'
true d3
=>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect
- x1)
- | None =>
- default
- end
- | _ =>
- default
- end
- | None =>
- default
- 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
- _ _ _ _ _
- _ _) _ =>
- default
- | _ =>
- default
- end
+ x2)
| None =>
- default0
- tt
+ UnderLets.Base
+ (x * x0)%expr
end
| _ =>
- default0
- tt
- end in
- match x1 with
- | #(idc0)%expr_pat =>
+ 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
+ (#(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
+ (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
- idc0
+ idc3
with
| @ident.Literal
- t1 v =>
+ t6 v1 =>
match
- t1 as t2
+ t6 as t7
return
(base.base_interp
- t2 ->
+ t7 ->
option Z)
with
| base.type.unit =>
@@ -2801,8 +12542,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v2 : Z =>
+ Some v2
| base.type.bool =>
fun
_ : bool
@@ -2811,29 +12552,1403 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
+ end v1
| _ =>
None
end
with
| Some
- args0 =>
+ args3 =>
match
- x0
+ s0 as t6
+ return
+ (Compile.value'
+ false t6 ->
+ UnderLets.UnderLets
+ base.type
+ ident var
+ (defaults.expr
+ (type.base
+ base.type.Z)))
with
- | #
- (idc1)%expr_pat =>
+ | 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
- idc1
+ 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
+ (#(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
+ (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
- t2 v =>
+ t6 v1 =>
match
- t2 as t3
+ t6 as t7
return
(base.base_interp
- t3 ->
+ t7 ->
option Z)
with
| base.type.unit =>
@@ -2842,8 +13957,8 @@ match idc in (ident t) return (Compile.value' true t) with
=> None
| base.type.Z =>
fun
- v0 : Z =>
- Some v0
+ v2 : Z =>
+ Some v2
| base.type.bool =>
fun
_ : bool
@@ -2852,18 +13967,66 @@ match idc in (ident t) return (Compile.value' true t) with
fun
_ : nat
=> None
- end v
+ 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
- args1 =>
+ x'1 =>
match
- s0 as t2
+ s3 as t7
return
(Compile.value'
- false t2 ->
+ false t7 ->
UnderLets.UnderLets
base.type
ident var
@@ -2872,19 +14035,19 @@ match idc in (ident t) return (Compile.value' true t) with
base.type.Z)))
with
| type.base
- t2 =>
+ t7 =>
fun
- v :
+ v2 :
defaults.expr
(type.base
- t2) =>
+ t7) =>
base.try_make_transport_cps
(fun
- t3 : base.type
+ t8 : base.type
=>
defaults.expr
(type.base
- t3)) t2
+ t8)) t7
base.type.Z
(UnderLets.UnderLets
base.type
@@ -2893,172 +14056,1413 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a :
+ a2 :
option
(defaults.expr
(type.base
- t2) ->
+ t7) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a
+ a2
with
| Some
- x' =>
+ x'2 =>
if
- args0 =?
+ (args0 =?
2
^
(2 *
Z.log2_up
args0 / 2) -
- 1
- then
- match
- invert_high
+ 1) &&
+ (args3 =?
+ 2
+ ^
(2 *
Z.log2_up
- args0)
- args1
- with
- | Some
- y =>
+ args0 / 2) -
+ 1)
+ then
UnderLets.Base
(#
- (ident.fancy_mullh
+ (ident.fancy_mulll
(2 *
Z.log2_up
args0))%expr @
(
- x' v,
- ##
- (y)%expr))%expr_pat
- | None =>
- default
- end
+ x'1 v1,
+ x'2 v2))%expr_pat
else
- default
+ UnderLets.Base
+ (x * x0)%expr
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
- | (s1 ->
- d1)%ptype =>
+ | (s4 ->
+ d4)%ptype =>
fun
_ :
Compile.value'
- false s1 ->
+ false s4 ->
Compile.value'
- true d1
+ true d4
=>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect
- x2)
+ x5)
| None =>
- default
- end
- | _ =>
- default
+ 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 =>
- default0
- tt
+ UnderLets.Base
+ (x * x0)%expr
end
| _ =>
- default0
- tt
- end in
- match x2 with
- | #(idc0)%expr_pat =>
+ 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
- match idc0 with
- | @ident.Literal t1 v =>
+ 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
- t1 as t2
- return
- (base.base_interp
- t2 ->
- option Z)
+ invert_high
+ (2 *
+ Z.log2_up args0)
+ args1
with
- | base.type.unit =>
- fun _ : unit =>
- None
- | base.type.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
+ | 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
+ (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
- t3 ->
+ t7 ->
option Z)
- with
- | base.type.unit =>
+ with
+ | base.type.unit =>
fun
_ : unit
=> None
- | base.type.Z =>
+ | base.type.Z =>
fun
- v0 : Z =>
- Some v0
- | base.type.bool =>
+ v2 : Z =>
+ Some v2
+ | base.type.bool =>
fun
_ : bool
=> None
- | base.type.nat =>
+ | base.type.nat =>
fun
_ : nat
=> None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s as t2
- return
- (Compile.value'
- false t2 ->
- UnderLets.UnderLets
+ 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
- t2 =>
- fun
- v :
+ 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
- t2) =>
- base.try_make_transport_cps
+ t7) =>
+ base.try_make_transport_cps
(fun
- t3 : base.type
+ t8 : base.type
=>
defaults.expr
(type.base
- t3)) t2
+ t8)) t7
base.type.Z
(UnderLets.UnderLets
base.type
@@ -3067,281 +15471,939 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base
base.type.Z)))
(fun
- a :
+ a2 :
option
(defaults.expr
(type.base
- t2) ->
+ t7) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a
+ a2
with
| Some
- x' =>
+ x'2 =>
if
- args0 =?
+ (args0 =?
2
^
(2 *
Z.log2_up
args0 / 2) -
- 1
- then
- match
- invert_high
+ 1) &&
+ (args3 =?
+ 2
+ ^
(2 *
Z.log2_up
- args0)
- args1
- with
- | Some
- y =>
+ args0 / 2) -
+ 1)
+ then
UnderLets.Base
(#
- (ident.fancy_mullh
+ (ident.fancy_mulll
(2 *
Z.log2_up
args0))%expr @
(
- x' v,
- ##
- (y)%expr))%expr_pat
+ x'1 v1,
+ x'2 v2))%expr_pat
+ else
+ UnderLets.Base
+ (x * x0)%expr
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
+ end)
+ | (s3 ->
+ d3)%ptype =>
+ fun
+ _ :
+ Compile.value'
+ false s3 ->
+ Compile.value'
+ true d3
+ =>
+ UnderLets.Base
+ (x * x0)%expr
end
- else
- default
+ (Compile.reflect
+ x5)
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
- | (s1 -> d1)%ptype =>
- fun
+ | (s3 ->
+ d3)%ptype =>
+ fun
_ :
Compile.value'
- false s1 ->
+ false s3 ->
Compile.value'
- true d1
- => default
- end
- (Compile.reflect
- x1)
- | None => default
- end
- | _ => default
- end
- | None => default0 tt
- end
- | _ => default0 tt
- end in
- 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
+ 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,
- ##(y)%expr))%expr_pat
- | None => default
- end
- else default
- | None => default
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false
- s1 ->
- Compile.value' true d1
- => default
- end (Compile.reflect x2)
- | None => default
- end
- | _ => default
- end
- | None => default0 tt
- end
- | _ => default0 tt
- end in
- 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 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
+ (
+ 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
- invert_low (2 * Z.log2_up args0)
- args1
+ 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 y =>
- UnderLets.Base
- (#(ident.fancy_mulll
- (2 * Z.log2_up args0))%expr @
- (x' v, ##(y)%expr))%expr_pat
- | None => default
+ | 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
- else default
- | None => default
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 => default
- end (Compile.reflect x1)
- | None => default
+ | _ => 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
- | _ => default
+ | None => UnderLets.Base (x * x0)%expr
end
- | None => default0 tt
+ | _ => UnderLets.Base (x * x0)%expr
end
- | _ => default0 tt
end
| None =>
match
@@ -3466,26 +16528,33 @@ match idc in (ident t) return (Compile.value' true t) with
args0))%expr @
(x'0 v0,
##(y)%expr))%expr_pat
- | None => default
+ | None =>
+ UnderLets.Base
+ (x * x0)%expr
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s1 -> d1)%ptype =>
fun
_ : Compile.value' false
s1 ->
Compile.value' true
- d1 => default
+ d1 =>
+ UnderLets.Base
+ (x * x0)%expr
end (Compile.reflect x2)
end
- | None => default
+ | None => UnderLets.Base (x * x0)%expr
end)
| (s1 -> d1)%ptype =>
fun
_ : Compile.value' false s1 ->
- Compile.value' true d1 => default
+ Compile.value' true d1 =>
+ UnderLets.Base (x * x0)%expr
end (Compile.reflect x2)
- | None => default
+ | None => UnderLets.Base (x * x0)%expr
end
| @expr.App _ _ _ s1 _
(@expr.App _ _ _ s2 _ #(idc1)%expr_pat x4) x3 =>
@@ -3496,179 +16565,6 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some _ =>
- llet default0 := fun 'tt =>
- 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
- default
- | None =>
- default
- end)
- | (s3 ->
- d3)%ptype =>
- fun
- _ :
- Compile.value'
- false s3 ->
- Compile.value'
- true d3
- =>
- default
- end
- (Compile.reflect
- x4)
- | None =>
- default
- end)
- | (s3 -> d3)%ptype =>
- fun
- _ : Compile.value'
- false s3 ->
- Compile.value'
- true d3
- => default
- end
- (Compile.reflect
- x2)
- | None => default
- end
- | _ => default
- end in
match x4 with
| #(idc2)%expr_pat =>
match
@@ -3774,8 +16670,207 @@ match idc in (ident t) return (Compile.value' true t) with
args0))%expr @
(x' v,
x'0 v0))%expr_pat
- else default
- | None => default
+ 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
@@ -3783,19 +16878,625 @@ match idc in (ident t) return (Compile.value' true t) with
false s3 ->
Compile.value'
true d3 =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end (Compile.reflect x3)
- | None => default
+ | None =>
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 -> d3)%ptype =>
fun
_ : Compile.value' false s3 ->
Compile.value' true d3 =>
- default
+ UnderLets.Base (x * x0)%expr
end (Compile.reflect x2)
- | None => default0 tt
+ | 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
- | _ => default0 tt
end
| None =>
match
@@ -3917,9 +17618,11 @@ match idc in (ident t) return (Compile.value' true t) with
x' v,
x'0 v0))%expr_pat
else
- default
+ UnderLets.Base
+ (x * x0)%expr
| None =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 -> d3)%ptype =>
fun
@@ -3927,22 +17630,27 @@ match idc in (ident t) return (Compile.value' true t) with
false s3 ->
Compile.value'
true d3 =>
- default
+ UnderLets.Base
+ (x * x0)%expr
end
(Compile.reflect x4)
- | None => default
+ | None =>
+ UnderLets.Base
+ (x * x0)%expr
end)
| (s3 -> d3)%ptype =>
fun
_ : Compile.value' false s3 ->
Compile.value' true d3
- => default
+ =>
+ UnderLets.Base (x * x0)%expr
end (Compile.reflect x2)
- | None => default
+ | None =>
+ UnderLets.Base (x * x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x * x0)%expr
end
- | None => default
+ | None => UnderLets.Base (x * x0)%expr
end
end
| @expr.App _ _ _ s1 _
@@ -3953,111 +17661,90 @@ match idc in (ident t) return (Compile.value' true t) with
(@expr.App _ _ _ s2 _ (_ @ _)%expr_pat _) _ |
@expr.App _ _ _ s1 _
(@expr.App _ _ _ s2 _ (@expr.LetIn _ _ _ _ _ _ _)
- _) _ => default
+ _) _ => UnderLets.Base (x * x0)%expr
| @expr.App _ _ _ s1 _ #(_)%expr_pat _ | @expr.App _
_ _ s1 _ ($_)%expr _ | @expr.App _ _ _ s1 _
(@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s1 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => default
- | _ => default
+ (@expr.LetIn _ _ _ _ _ _ _) _ =>
+ UnderLets.Base (x * x0)%expr
+ | _ => UnderLets.Base (x * x0)%expr
end
- | None => default
+ | None => UnderLets.Base (x * x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x * x0)%expr
end
- | None => default
+ | 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 _ _ _ _ _ _ _) _) _ => default
+ 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 _ _ _ _ _ _ _) _ => default
- | _ => default
+ (@expr.LetIn _ _ _ _ _ _ _) _ => UnderLets.Base (x * x0)%expr
+ | _ => UnderLets.Base (x * x0)%expr
end
| ident.Z_pow =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat
| ident.Z_sub =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x - x0)%expr in
- default
+ UnderLets.Base (x - x0)%expr
| ident.Z_opp =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (- x)%expr in
- default
+ UnderLets.Base (- x)%expr
| ident.Z_div =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x / x0)%expr in
- default
+ UnderLets.Base (x / x0)%expr
| ident.Z_modulo =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x mod x0)%expr in
- default
+ UnderLets.Base (x mod x0)%expr
| ident.Z_log2 =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_log2)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_log2)%expr @ x)%expr_pat
| ident.Z_log2_up =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_log2_up)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_log2_up)%expr @ x)%expr_pat
| ident.Z_eqb =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat
| ident.Z_leb =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat
| ident.Z_geb =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat
| ident.Z_of_nat =>
fun x : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Z_of_nat)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_of_nat)%expr @ x)%expr_pat
| ident.Z_to_nat =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_to_nat)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_to_nat)%expr @ x)%expr_pat
| ident.Z_shiftr =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x >> x0)%expr in
- default
+ UnderLets.Base (x >> x0)%expr
| ident.Z_shiftl =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x << x0)%expr in
- default
+ UnderLets.Base (x << x0)%expr
| ident.Z_land =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x &' x0)%expr in
- default
+ UnderLets.Base (x &' x0)%expr
| ident.Z_lor =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x || x0)%expr in
- default
+ UnderLets.Base (x || x0)%expr
| ident.Z_bneg =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_bneg)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_bneg)%expr @ x)%expr_pat
| ident.Z_lnot_modulo =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_lnot_modulo)%expr @ x @ x0)%expr_pat in
- default
+ 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) =>
- llet default := UnderLets.Base
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat in
- default
+ 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) =>
- llet default := UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -4073,130 +17760,4102 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args =>
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- llet default0 :=
- fun 'tt =>
- if
- args =?
- 2
- ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_add
+ 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 default in
- match x0 with
- | (@expr.App _
- _ _ s0 _
- #(idc0)
- x3 @ x2)%expr_pat =>
- 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
+ (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
- v0 : Z =>
- Some v0
- | base.type.bool =>
- fun
+ => None
+ | base.type.Z =>
+ fun v1 : Z
+ =>
+ Some v1
+ | 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
+ => 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)))
- with
- | type.base
- t2 =>
- fun
- v :
- defaults.expr
+ (fun
+ a0 :
+ option
+ (defaults.expr
(type.base
- t2) =>
- base.try_make_transport_cps
- (fun
- t3 : base.type
- =>
+ t5) ->
defaults.expr
(type.base
- t3)) t2
- base.type.Z
- (UnderLets.UnderLets
+ 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
- a :
+ (fun
+ a0 :
option
(defaults.expr
(type.base
- t2) ->
+ t5) ->
defaults.expr
(type.base
base.type.Z))
- =>
- match
- a
- with
- | Some
- x' =>
+ =>
+ match
+ a0
+ with
+ | Some
+ x'0 =>
if
args =?
2
@@ -4209,141 +21868,723 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.fancy_add
(Z.log2
args)
- (- args1))%expr @
- (x1,
- x' v))%expr_pat
+ (- args3))%expr @
+ (x0,
+ x'0 v0))%expr_pat
else
- default
- | None =>
- default
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- default
- end
- (Compile.reflect
- x3)
- | None =>
- default
- end
- | _ =>
- default
- 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
| None =>
- default
- 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 =>
- default
- | (#(_) @ _)%expr_pat |
- (($_)%expr @
- _)%expr_pat |
- (@expr.Abs _
- _ _ _ _ _ @
- _)%expr_pat |
- (@expr.LetIn
- _ _ _ _ _ _
- _ @ _)%expr_pat =>
- default
- | _ =>
- default0
- tt
- end in
- match x1 with
- | (@expr.App _ _ _ s0 _ #
- (idc0) x3 @ x2)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftr =>
- Some tt
- | _ => 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
- | Some _ =>
- match x2 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal
- t2 v =>
- match
- t2 as t3
- return
+ | base.type.unit =>
+ fun _ : unit => None
+ | base.type.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
- t3 ->
+ t6 ->
option Z)
- with
- | base.type.unit =>
+ with
+ | base.type.unit =>
fun
_ : unit
=> None
- | base.type.Z =>
+ | base.type.Z =>
fun
- v0 : Z =>
- Some v0
- | base.type.bool =>
+ v1 : Z =>
+ Some v1
+ | base.type.bool =>
fun
_ : bool
=> None
- | base.type.nat =>
+ | base.type.nat =>
fun
_ : nat
=> None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value'
- false t2 ->
- UnderLets.UnderLets
+ 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
- t2 =>
- fun
- v :
+ with
+ | type.base
+ t5 =>
+ fun
+ v0 :
defaults.expr
(type.base
- t2) =>
- base.try_make_transport_cps
+ t5) =>
+ base.try_make_transport_cps
(fun
- t3 : base.type
+ t6 : base.type
=>
defaults.expr
(type.base
- t3)) t2
+ t6)) t5
base.type.Z
(UnderLets.UnderLets
base.type
@@ -4353,20 +22594,20 @@ match idc in (ident t) return (Compile.value' true t) with
(base.type.Z *
base.type.Z)%etype)))
(fun
- a :
+ a0 :
option
(defaults.expr
(type.base
- t2) ->
+ t5) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a
+ a0
with
| Some
- x' =>
+ x'0 =>
if
args =?
2
@@ -4379,117 +22620,690 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.fancy_add
(Z.log2
args)
- (- args1))%expr @
- (x0,
- x' v))%expr_pat
+ (- 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
- default
+ UnderLets.Base
+ (#
+ (ident.Z_add_get_carry)%expr @
+ x @ x0 @
+ x1)%expr_pat
| None =>
- default
+ UnderLets.Base
+ (#
+ (ident.Z_add_get_carry)%expr @
+ x @ x0 @
+ x1)%expr_pat
end)
- | (s1 -> d1)%ptype =>
- fun
+ | (s3 -> d3)%ptype =>
+ fun
_ :
Compile.value'
- false s1 ->
+ false s3 ->
Compile.value'
- true d1
- => default
- end
- (Compile.reflect
- x3)
- | None => default
- end
- | _ => default
- end
- | None => default
+ 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
- | (@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 => default
- | (#(_) @ _)%expr_pat |
- (($_)%expr @ _)%expr_pat |
- (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat |
- (@expr.LetIn _ _ _ _ _ _ _ @
- _)%expr_pat => default
- | _ => default0 tt
- end in
- 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 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))
+ 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)
=>
- match a with
- | Some x' =>
+ 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
@@ -4497,40 +23311,3925 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base
(#(ident.fancy_add
(Z.log2 args)
- args1)%expr @
- (x1, x' v))%expr_pat
- else default
- | None => default
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false
- s1 ->
- Compile.value' true d1
- => default
- end (Compile.reflect x3)
- | None => default
- end
- | _ => default
- end
- | None => default
- 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 =>
- default
- | (#(_) @ _)%expr_pat |
- (($_)%expr @ _)%expr_pat |
- (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat |
- (@expr.LetIn _ _ _ _ _ _ _ @ _)%expr_pat =>
- default
- | _ => default0 tt
- end in
- match x1 with
- | (@expr.App _ _ _ s0 _ #(idc0) x3 @ x2)%expr_pat =>
+ 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
+| 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
+ | _ => 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
@@ -4538,7 +27237,7 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some _ =>
- match x2 with
+ match x3 with
| #(idc1)%expr_pat =>
match
match idc1 with
@@ -4585,185 +27284,2560 @@ match idc in (ident t) return (Compile.value' true t) with
if args =? 2 ^ Z.log2 args
then
UnderLets.Base
- (#(ident.fancy_add (Z.log2 args)
- args1)%expr @ (x0, x' v))%expr_pat
- else default
- | None => default
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 => default
- end (Compile.reflect x3)
- | None => default
- end
- | _ => default
- end
- | None => default
- 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 =>
- default
- | (#(_) @ _)%expr_pat | (($_)%expr @ _)%expr_pat |
- (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat |
- (@expr.LetIn _ _ _ _ _ _ _ @ _)%expr_pat => default
- | _ => default0 tt
- end
- | None => default
- end
- | _ => default
- end
-| ident.Z_add_with_carry =>
- fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat in
- default
-| ident.Z_add_with_get_carry =>
- fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat in
- 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 =>
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- llet default0 := fun 'tt =>
- llet default0 :=
- fun 'tt =>
- if
- args =?
- 2
- ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_addc
+ (#(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 default in
- match x1 with
- | (@expr.App _
- _ _ s0 _
- #(idc0)
- x4 @ x3)%expr_pat =>
- 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
+ (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
- v0 : Z =>
- Some v0
- | base.type.bool =>
- fun
+ => None
+ | base.type.Z =>
+ fun v1 : Z
+ =>
+ Some v1
+ | 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
+ => 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)))
- with
- | type.base
- t2 =>
- fun
- v :
- defaults.expr
+ (fun
+ a0 :
+ option
+ (defaults.expr
(type.base
- t2) =>
- base.try_make_transport_cps
- (fun
- t3 : base.type
- =>
+ t5) ->
defaults.expr
(type.base
- t3)) t2
- base.type.Z
- (UnderLets.UnderLets
+ 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
- a :
+ (fun
+ a0 :
option
(defaults.expr
(type.base
- t2) ->
+ t5) ->
defaults.expr
(type.base
base.type.Z))
- =>
- match
- a
- with
- | Some
- x' =>
+ =>
+ match
+ a0
+ with
+ | Some
+ x'0 =>
if
args =?
2
@@ -4776,141 +29850,727 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.fancy_addc
(Z.log2
args)
- (- args1))%expr @
- (x0, x2,
- x' v))%expr_pat
+ (- args3))%expr @
+ (x0, x1,
+ x'0 v0))%expr_pat
else
- default
- | None =>
- default
- end)
- | (s1 ->
- d1)%ptype =>
- fun
- _ :
- Compile.value'
- false s1 ->
- Compile.value'
- true d1
- =>
- default
- end
- (Compile.reflect
- x4)
- | None =>
- default
- end
- | _ =>
- default
- 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
| None =>
- default
- 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 =>
- default
- | (#(_) @ _)%expr_pat |
- (($_)%expr @
- _)%expr_pat |
- (@expr.Abs _
- _ _ _ _ _ @
- _)%expr_pat |
- (@expr.LetIn
- _ _ _ _ _ _
- _ @ _)%expr_pat =>
- default
- | _ =>
- default0
- tt
- end in
- match x2 with
- | (@expr.App _ _ _ s0 _ #
- (idc0) x4 @ x3)%expr_pat =>
- match
- match idc0 with
- | ident.Z_shiftr =>
- Some tt
- | _ => 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
- | Some _ =>
- match x3 with
- | #(idc1)%expr_pat =>
- match
- match idc1 with
- | @ident.Literal
- t2 v =>
- match
- t2 as t3
- return
+ | base.type.unit =>
+ fun _ : unit => None
+ | base.type.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
- t3 ->
+ t6 ->
option Z)
- with
- | base.type.unit =>
+ with
+ | base.type.unit =>
fun
_ : unit
=> None
- | base.type.Z =>
+ | base.type.Z =>
fun
- v0 : Z =>
- Some v0
- | base.type.bool =>
+ v1 : Z =>
+ Some v1
+ | base.type.bool =>
fun
_ : bool
=> None
- | base.type.nat =>
+ | base.type.nat =>
fun
_ : nat
=> None
- end v
- | _ => None
- end
- with
- | Some args1 =>
- match
- s0 as t2
- return
- (Compile.value'
- false t2 ->
- UnderLets.UnderLets
+ 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
- t2 =>
- fun
- v :
+ with
+ | type.base
+ t5 =>
+ fun
+ v0 :
defaults.expr
(type.base
- t2) =>
- base.try_make_transport_cps
+ t5) =>
+ base.try_make_transport_cps
(fun
- t3 : base.type
+ t6 : base.type
=>
defaults.expr
(type.base
- t3)) t2
+ t6)) t5
base.type.Z
(UnderLets.UnderLets
base.type
@@ -4920,20 +30580,20 @@ match idc in (ident t) return (Compile.value' true t) with
(base.type.Z *
base.type.Z)%etype)))
(fun
- a :
+ a0 :
option
(defaults.expr
(type.base
- t2) ->
+ t5) ->
defaults.expr
(type.base
base.type.Z))
=>
match
- a
+ a0
with
| Some
- x' =>
+ x'0 =>
if
args =?
2
@@ -4946,117 +30606,703 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.fancy_addc
(Z.log2
args)
- (- args1))%expr @
+ (- 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,
- x' v))%expr_pat
+ x2))%expr_pat
else
- default
+ UnderLets.Base
+ (#
+ (ident.Z_add_with_get_carry)%expr @
+ x @ x0 @
+ x1 @ x2)%expr_pat
| None =>
- default
+ UnderLets.Base
+ (#
+ (ident.Z_add_with_get_carry)%expr @
+ x @ x0 @
+ x1 @ x2)%expr_pat
end)
- | (s1 -> d1)%ptype =>
- fun
+ | (s3 -> d3)%ptype =>
+ fun
_ :
Compile.value'
- false s1 ->
+ false s3 ->
Compile.value'
- true d1
- => default
- end
- (Compile.reflect
- x4)
- | None => default
- end
- | _ => default
- end
- | None => default
+ 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
- | (@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 => default
- | (#(_) @ _)%expr_pat |
- (($_)%expr @ _)%expr_pat |
- (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat |
- (@expr.LetIn _ _ _ _ _ _ _ @
- _)%expr_pat => default
- | _ => default0 tt
- end in
- 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 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))
+ 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)
=>
- match a with
- | Some x' =>
+ 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
@@ -5064,128 +31310,2408 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base
(#(ident.fancy_addc
(Z.log2 args)
- args1)%expr @
- (x0, x2, x' v))%expr_pat
- else default
- | None => default
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false
- s1 ->
- Compile.value' true d1
- => default
- end (Compile.reflect x4)
- | None => default
- end
- | _ => default
- end
- | None => default
- 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 =>
- default
- | (#(_) @ _)%expr_pat |
- (($_)%expr @ _)%expr_pat |
- (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat |
- (@expr.LetIn _ _ _ _ _ _ _ @ _)%expr_pat =>
- default
- | _ => default0 tt
- end in
- 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 =>
+ 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 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
+ match idc0 with
+ | ident.Z_shiftr => Some tt
| _ => None
end
with
- | Some args1 =>
+ | 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
- 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)))
+ 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
- | 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
+ | 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.fancy_addc (Z.log2 args)
- args1)%expr @ (x0, x1, x' v))%expr_pat
- else default
- | None => default
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 => default
- end (Compile.reflect x4)
- | None => default
+ (#(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
- | _ => default
end
- | None => default
+ | (@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
- | (@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 =>
- default
- | (#(_) @ _)%expr_pat | (($_)%expr @ _)%expr_pat |
- (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat |
- (@expr.LetIn _ _ _ _ _ _ _ @ _)%expr_pat => default
- | _ => default0 tt
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
| ident.Z_sub_get_borrow =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -5201,13 +33727,6 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args =>
- llet default0 := fun 'tt =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_sub (Z.log2 args) 0)%expr @
- (x0, x1))%expr_pat
- else default in
match x1 with
| (@expr.App _ _ _ s0 _ #(idc0) x3 @ x2)%expr_pat =>
match
@@ -5266,17 +33785,48 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base
(#(ident.fancy_sub (Z.log2 args)
args1)%expr @ (x0, x' v))%expr_pat
- else default
- | None => default
+ 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 => default
+ Compile.value' true d1 =>
+ UnderLets.Base
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @
+ x1)%expr_pat
end (Compile.reflect x3)
- | None => default
+ | 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
- | _ => default
+ | _ =>
+ 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
@@ -5337,39 +33887,92 @@ match idc in (ident t) return (Compile.value' true t) with
(#(ident.fancy_sub
(Z.log2 args) (- args1))%expr @
(x0, x' v))%expr_pat
- else default
- | None => default
+ 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 => default
+ Compile.value' true d1 =>
+ UnderLets.Base
+ (#(ident.Z_sub_get_borrow)%expr @ x @
+ x0 @ x1)%expr_pat
end (Compile.reflect x3)
- | None => default
+ | 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
- | _ => default
+ | _ =>
+ 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 => default
+ | 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 =>
- default
- | (#(_) @ _)%expr_pat | (($_)%expr @ _)%expr_pat |
- (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat |
- (@expr.LetIn _ _ _ _ _ _ _ @ _)%expr_pat => default
- | _ => default0 tt
+ 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 => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
end
| ident.Z_sub_with_get_borrow =>
fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -5385,13 +33988,6 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args =>
- llet default0 := fun 'tt =>
- if args =? 2 ^ Z.log2 args
- then
- UnderLets.Base
- (#(ident.fancy_subb (Z.log2 args) 0)%expr @
- (x0, x1, x2))%expr_pat
- else default in
match x2 with
| (@expr.App _ _ _ s0 _ #(idc0) x4 @ x3)%expr_pat =>
match
@@ -5450,17 +34046,50 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base
(#(ident.fancy_subb (Z.log2 args)
args1)%expr @ (x0, x1, x' v))%expr_pat
- else default
- | None => default
+ 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 => default
+ Compile.value' true d1 =>
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @
+ x0 @ x1 @ x2)%expr_pat
end (Compile.reflect x4)
- | None => default
+ | 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
- | _ => default
+ | _ =>
+ 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
@@ -5521,42 +34150,94 @@ match idc in (ident t) return (Compile.value' true t) with
(#(ident.fancy_subb
(Z.log2 args) (- args1))%expr @
(x0, x1, x' v))%expr_pat
- else default
- | None => default
+ 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 => default
+ Compile.value' true d1 =>
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @
+ x @ x0 @ x1 @ x2)%expr_pat
end (Compile.reflect x4)
- | None => default
+ | 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
- | _ => default
+ | _ =>
+ 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 => default
+ | 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 =>
- default
- | (#(_) @ _)%expr_pat | (($_)%expr @ _)%expr_pat |
- (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat |
- (@expr.LetIn _ _ _ _ _ _ _ @ _)%expr_pat => default
- | _ => default0 tt
+ 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 => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
| ident.Z_zselect =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat in
- llet default0 := fun 'tt =>
- UnderLets.Base
- (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat in
match x with
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x3) x2 =>
match match idc with
@@ -5606,17 +34287,27 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base
(#(ident.fancy_selm (Z.log2 args0))%expr @
(x' v, x0, x1))%expr_pat
- else default
- | None => default
+ 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 => default
+ Compile.value' true d1 =>
+ UnderLets.Base
+ (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
end (Compile.reflect x2)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_selc)%expr @ (x, x0, x1))%expr_pat
end
| None =>
match
@@ -5626,80 +34317,6 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some _ =>
- llet default1 := fun 'tt =>
- 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 default
- | None => default
- end)
- | (s1 -> d1)%ptype =>
- fun
- _ : Compile.value' false s1 ->
- Compile.value' true d1 =>
- default
- end (Compile.reflect x3)
- | None => default
- end
- | _ => default
- end in
match x3 with
| #(idc0)%expr_pat =>
match
@@ -5743,39 +34360,413 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base
(#(ident.fancy_sell)%expr @
(x' v, x0, x1))%expr_pat
- else default
- | None => default
+ 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 => default
+ Compile.value' true d1 =>
+ UnderLets.Base
+ (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
end (Compile.reflect x2)
- | None => default1 tt
+ | 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
- | _ => default1 tt
end
- | None => default
+ | 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 _ _ _ _ _ _ _) _) _ => default
+ 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 _ _ _ _ _ _ _) _ => default
- | _ => default0 tt
+ (@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
| ident.Z_add_modulo =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet _ := UnderLets.Base
- (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat in
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) =>
- llet default := UnderLets.Base
- (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -5813,100 +34804,81 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base
(#(ident.fancy_rshi (Z.log2 args) args0)%expr @
(x0, x1))%expr_pat
- else default
- | None => default
+ 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
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
| ident.Z_cc_m =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat
| ident.Z_cast range =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_cast range)%expr @ x)%expr_pat in
- default
+ 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) =>
- llet default := UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
| ident.fancy_add log2wordmax imm =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
| ident.fancy_addc log2wordmax imm =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
| ident.fancy_sub log2wordmax imm =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
| ident.fancy_subb log2wordmax imm =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
| ident.fancy_mulll log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
| ident.fancy_mullh log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
| ident.fancy_mulhl log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
| ident.fancy_mulhh log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
| ident.fancy_rshi log2wordmax x =>
fun x0 : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
| ident.fancy_selc =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat
| ident.fancy_selm log2wordmax =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
| ident.fancy_sell =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat
| ident.fancy_addm =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat in
- default
+ UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat
end
: Compile.value' true t
diff --git a/src/Experiments/NewPipeline/nbe_rewrite_head.out b/src/Experiments/NewPipeline/nbe_rewrite_head.out
index f64ef126e..82d4983bd 100644
--- a/src/Experiments/NewPipeline/nbe_rewrite_head.out
+++ b/src/Experiments/NewPipeline/nbe_rewrite_head.out
@@ -8,26 +8,13 @@ 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 =>
- llet _ := UnderLets.Base ##(v0)%expr in
- UnderLets.Base ##(v0)%expr
- | base.type.Z =>
- fun v0 : Z =>
- llet _ := UnderLets.Base ##(v0)%expr in
- UnderLets.Base ##(v0)%expr
- | base.type.bool =>
- fun v0 : bool =>
- llet _ := UnderLets.Base ##(v0)%expr in
- UnderLets.Base ##(v0)%expr
- | base.type.nat =>
- fun v0 : nat =>
- llet _ := UnderLets.Base ##(v0)%expr in
- 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) =>
- llet default := UnderLets.Base (#(ident.Nat_succ)%expr @ x)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -43,13 +30,12 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args => UnderLets.Base ##(Nat.succ args)%expr
- | None => default
+ | None => UnderLets.Base (#(ident.Nat_succ)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Nat_succ)%expr @ x)%expr_pat
end
| ident.Nat_pred =>
fun x : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_pred)%expr @ x)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -65,13 +51,12 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args => UnderLets.Base ##(Nat.pred args)%expr
- | None => default
+ | None => UnderLets.Base (#(ident.Nat_pred)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Nat_pred)%expr @ x)%expr_pat
end
| ident.Nat_max =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -104,17 +89,17 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##(Nat.max args args0)%expr
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Nat_max)%expr @ x @ x0)%expr_pat
end
| ident.Nat_mul =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -147,17 +132,17 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##((args * args0)%nat)%expr
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Nat_mul)%expr @ x @ x0)%expr_pat
end
| ident.Nat_add =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -190,17 +175,17 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##((args + args0)%nat)%expr
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Nat_add)%expr @ x @ x0)%expr_pat
end
| ident.Nat_sub =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -233,28 +218,25 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##((args - args0)%nat)%expr
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Nat_sub)%expr @ x @ x0)%expr_pat
end
-| @ident.nil t => llet default := UnderLets.Base []%expr_pat in
- default
+| @ident.nil t => UnderLets.Base []%expr_pat
| @ident.cons t =>
fun (x : defaults.expr (type.base t))
(x0 : defaults.expr (type.base (base.type.list t))) =>
- llet default := UnderLets.Base (x :: x0)%expr_pat in
- default
+ UnderLets.Base (x :: x0)%expr_pat
| @ident.pair A B =>
fun (x : defaults.expr (type.base A)) (x0 : defaults.expr (type.base B))
- => llet default := UnderLets.Base (x, x0)%expr_pat in
- default
+ => UnderLets.Base (x, x0)%expr_pat
| @ident.fst A B =>
fun x : defaults.expr (type.base (A * B)%etype) =>
- llet default := UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat in
match x with
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x1) x0 =>
match
@@ -293,30 +275,32 @@ match idc in (ident t) return (Compile.value' true t) with
defaults.expr (type.base A)) =>
match a with
| Some x' => UnderLets.Base (x' v)
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
end)
| (s1 -> d1)%ptype =>
fun _ : Compile.value' false s1 -> Compile.value' true d1
- => default
+ => UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
end (Compile.reflect x0)
| (s1 -> d1)%ptype =>
fun _ : Compile.value' false s1 -> Compile.value' true d1 =>
- default
+ UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
end (Compile.reflect x1)
- | None => default
+ | 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 _ _ _ _ _ _ _) _) _ => default
+ 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 _ _ _ _ _ _ _) _ => default
- | _ => default
+ (@expr.LetIn _ _ _ _ _ _ _) _ =>
+ UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
+ | _ => UnderLets.Base (#(ident.fst)%expr @ x)%expr_pat
end
| @ident.snd A B =>
fun x : defaults.expr (type.base (A * B)%etype) =>
- llet default := UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat in
match x with
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(idc)%expr_pat x1) x0 =>
match
@@ -355,26 +339,29 @@ match idc in (ident t) return (Compile.value' true t) with
defaults.expr (type.base B)) =>
match a with
| Some x' => UnderLets.Base (x' v0)
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
end)
| (s1 -> d1)%ptype =>
fun _ : Compile.value' false s1 -> Compile.value' true d1
- => default
+ => UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
end (Compile.reflect x0)
| (s1 -> d1)%ptype =>
fun _ : Compile.value' false s1 -> Compile.value' true d1 =>
- default
+ UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
end (Compile.reflect x1)
- | None => default
+ | 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 _ _ _ _ _ _ _) _) _ => default
+ 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 _ _ _ _ _ _ _) _ => default
- | _ => default
+ (@expr.LetIn _ _ _ _ _ _ _) _ =>
+ UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
+ | _ => UnderLets.Base (#(ident.snd)%expr @ x)%expr_pat
end
| @ident.prod_rect A B T =>
fun
@@ -383,11 +370,13 @@ 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)) =>
- llet default := 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 in
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
@@ -462,30 +451,80 @@ match idc in (ident t) return (Compile.value' true t) with
anyexpr_ty))) :=
fv in
unwrap))
- | None => default
+ | 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 => default
+ | 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 => default
+ | 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
- => default
+ =>
+ 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 =>
- default
+ 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 => default
+ | 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 _ _ _ _ _ _) _) _ | @expr.App
- _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ | @expr.App _ _ _
- s _ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ => default
- | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr _ |
- @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => default
- | _ => default
+ _ _ 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
| @ident.bool_rect T =>
fun
@@ -494,12 +533,6 @@ 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)) =>
- llet default := 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 in
match x1 with
| #(idc)%expr_pat =>
match
@@ -547,11 +580,36 @@ match idc in (ident t) return (Compile.value' true t) with
(let (anyexpr_ty, _) := a0 in anyexpr_ty))) :=
fv in
unwrap))
- | None => default
+ | 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))%under_lets
- | None => default
+ | 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
- | _ => default
+ | ($_)%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
| @ident.nat_rect P =>
fun
@@ -563,14 +621,6 @@ 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)) =>
- llet default := 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 in
match x1 with
| #(idc)%expr_pat =>
match
@@ -632,13 +682,47 @@ match idc in (ident t) return (Compile.value' true t) with
(let (anyexpr_ty, _) := a0 in
anyexpr_ty))) := fv in
unwrap))
- | None => default
+ | 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 => default
+ | 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 => default
+ | 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
- | _ => default
+ | ($_)%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
| @ident.nat_rect_arrow P Q =>
fun
@@ -654,21 +738,6 @@ 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)) =>
- llet default := 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 in
match x1 with
| #(idc)%expr_pat =>
match
@@ -750,15 +819,100 @@ match idc in (ident t) return (Compile.value' true t) with
a0 in
anyexpr_ty))) := fv in
unwrap))
- | None => default
+ | 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))%under_lets
- | None => default
+ | 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 => default
+ | 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 => default
+ | 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
- | _ => default
+ | ($_)%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
| @ident.list_rect A P =>
fun
@@ -771,18 +925,6 @@ 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))) =>
- llet default := 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 in
Compile.castv x0
(UnderLets.UnderLets base.type ident var (defaults.expr (type.base P)))
(fun
@@ -838,11 +980,40 @@ match idc in (ident t) return (Compile.value' true t) with
(let (anyexpr_ty, _) := a0 in
anyexpr_ty))) := fv in
unwrap))
- | None => default
+ | 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))%under_lets
- | None => default
+ | 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 => default
+ | 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)
| @ident.list_case A P =>
fun
@@ -854,15 +1025,6 @@ 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))) =>
- llet default := 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 in
match x1 with
| #(idc)%expr_pat =>
match match idc with
@@ -895,10 +1057,36 @@ match idc in (ident t) return (Compile.value' true t) with
(let (anyexpr_ty, _) := a0 in anyexpr_ty))) :=
fv in
unwrap))
- | None => default
+ | 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))%under_lets
- | None => default
+ | 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
@@ -974,43 +1162,132 @@ match idc in (ident t) return (Compile.value' true t) with
anyexpr_ty))) :=
fv in
unwrap))
- | None => default
+ | 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 => default
+ | 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 => default
+ | 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
- => default
+ =>
+ 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 x2)
| (s1 -> d1)%ptype =>
fun _ : Compile.value' false s1 -> Compile.value' true d1 =>
- default
+ 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 => default
+ | 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 _ _ _ _ _ _) _) _ | @expr.App
- _ _ _ s _ (@expr.App _ _ _ s0 _ (_ @ _)%expr_pat _) _ | @expr.App _ _ _
- s _ (@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ => default
- | @expr.App _ _ _ s _ #(_)%expr_pat _ | @expr.App _ _ _ s _ ($_)%expr _ |
- @expr.App _ _ _ s _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => default
- | _ => default
+ _ _ 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
| @ident.List_length T =>
fun x : defaults.expr (type.base (base.type.list T)) =>
- llet default := UnderLets.Base (#(ident.List_length)%expr @ x)%expr_pat in
reflect_list_cps x
(fun xs : option (list (defaults.expr (type.base T))) =>
match xs with
| Some xs0 => UnderLets.Base ##(length xs0)%expr
- | None => default
+ | None => UnderLets.Base (#(ident.List_length)%expr @ x)%expr_pat
end)
| ident.List_seq =>
fun x x0 : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -1050,19 +1327,18 @@ match idc in (ident t) return (Compile.value' true t) with
(type.base (base.type.list base.type.nat)))
=> (x1 :: xs)%expr_pat) []%expr_pat
(map (fun v : nat => ##(v)%expr) (seq args args0)))
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.List_seq)%expr @ x @ x0)%expr_pat
end
| @ident.List_firstn A =>
fun (x : defaults.expr (type.base base.type.nat))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- llet default := UnderLets.Base
- (#(ident.List_firstn)%expr @ x @ x0)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -1095,19 +1371,22 @@ match idc in (ident t) return (Compile.value' true t) with
match a with
| Some x' =>
UnderLets.Base (x' (reify_list (firstn args xs0)))
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_firstn)%expr @ x @ x0)%expr_pat
end)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_firstn)%expr @ x @ x0)%expr_pat
end)
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.List_firstn)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.List_firstn)%expr @ x @ x0)%expr_pat
end
| @ident.List_skipn A =>
fun (x : defaults.expr (type.base base.type.nat))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- llet default := UnderLets.Base
- (#(ident.List_skipn)%expr @ x @ x0)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -1140,19 +1419,21 @@ match idc in (ident t) return (Compile.value' true t) with
match a with
| Some x' =>
UnderLets.Base (x' (reify_list (skipn args xs0)))
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_skipn)%expr @ x @ x0)%expr_pat
end)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_skipn)%expr @ x @ x0)%expr_pat
end)
- | None => default
+ | None => UnderLets.Base (#(ident.List_skipn)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.List_skipn)%expr @ x @ x0)%expr_pat
end
| @ident.List_repeat A =>
fun (x : defaults.expr (type.base A))
(x0 : defaults.expr (type.base base.type.nat)) =>
- llet default := UnderLets.Base
- (#(ident.List_repeat)%expr @ x @ x0)%expr_pat in
match x0 with
| #(idc)%expr_pat =>
match
@@ -1179,17 +1460,18 @@ match idc in (ident t) return (Compile.value' true t) with
defaults.expr (type.base (base.type.list A))) =>
match a with
| Some x' => UnderLets.Base (x' (reify_list (repeat x args)))
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_repeat)%expr @ x @ x0)%expr_pat
end)
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.List_repeat)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.List_repeat)%expr @ x @ x0)%expr_pat
end
| @ident.List_combine A B =>
fun (x : defaults.expr (type.base (base.type.list A)))
(x0 : defaults.expr (type.base (base.type.list B))) =>
- llet default := UnderLets.Base
- (#(ident.List_combine)%expr @ x @ x0)%expr_pat in
reflect_list_cps x
(fun xs : option (list (defaults.expr (type.base A))) =>
match xs with
@@ -1226,11 +1508,16 @@ match idc in (ident t) return (Compile.value' true t) with
(reify_list
(map (fun '(x1, y) => (x1, y)%expr_pat)
(combine xs0 ys0))))
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_combine)%expr @ x @ x0)%expr_pat
end)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_combine)%expr @ x @ x0)%expr_pat
end)
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.List_combine)%expr @ x @ x0)%expr_pat
end)
| @ident.List_map A B =>
fun
@@ -1238,10 +1525,6 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base B)))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- llet default := UnderLets.Base
- (#(ident.List_map)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat in
Compile.castbe x0
(UnderLets.UnderLets base.type ident var
(defaults.expr (type.base (base.type.list B))))
@@ -1302,15 +1585,26 @@ match idc in (ident t) return (Compile.value' true t) with
(let (anyexpr_ty, _) := a0 in
anyexpr_ty))) := fv in
unwrap))
- | None => default
+ | 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 => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_map)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
end)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_map)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
end)
| @ident.List_app A =>
fun x x0 : defaults.expr (type.base (base.type.list A)) =>
- llet default := UnderLets.Base (x ++ x0)%expr in
Compile.castbe x
(UnderLets.UnderLets base.type ident var
(defaults.expr (type.base (base.type.list A))))
@@ -1370,15 +1664,14 @@ match idc in (ident t) return (Compile.value' true t) with
(let (anyexpr_ty, _) := a0 in
anyexpr_ty))) := fv in
unwrap))
- | None => default
+ | None => UnderLets.Base (x ++ x0)%expr
end))%under_lets
- | None => default
+ | None => UnderLets.Base (x ++ x0)%expr
end)
- | None => default
+ | None => UnderLets.Base (x ++ x0)%expr
end)
| @ident.List_rev A =>
fun x : defaults.expr (type.base (base.type.list A)) =>
- llet default := UnderLets.Base (#(ident.List_rev)%expr @ x)%expr_pat in
reflect_list_cps x
(fun xs : option (list (defaults.expr (type.base A))) =>
match xs with
@@ -1394,9 +1687,9 @@ match idc in (ident t) return (Compile.value' true t) with
defaults.expr (type.base (base.type.list A))) =>
match a with
| Some x' => UnderLets.Base (x' (reify_list (rev xs0)))
- | None => default
+ | None => UnderLets.Base (#(ident.List_rev)%expr @ x)%expr_pat
end)
- | None => default
+ | None => UnderLets.Base (#(ident.List_rev)%expr @ x)%expr_pat
end)
| @ident.List_flat_map A B =>
fun
@@ -1404,10 +1697,6 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base (base.type.list B))))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- llet default := UnderLets.Base
- (#(ident.List_flat_map)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat in
Compile.castbe x0
(UnderLets.UnderLets base.type ident var
(defaults.expr (type.base (base.type.list B))))
@@ -1466,11 +1755,23 @@ match idc in (ident t) return (Compile.value' true t) with
=>
match a with
| Some x' => UnderLets.Base (x' fv0)
- | None => default
+ | 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 => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_flat_map)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
end)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_flat_map)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
end)
| @ident.List_partition A =>
fun
@@ -1478,10 +1779,6 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.UnderLets base.type ident var
(defaults.expr (type.base base.type.bool)))
(x0 : defaults.expr (type.base (base.type.list A))) =>
- llet default := UnderLets.Base
- (#(ident.List_partition)%expr @
- (λ x1 : var (type.base A),
- UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat in
Compile.castbe x0
(UnderLets.UnderLets base.type ident var
(defaults.expr
@@ -1562,11 +1859,23 @@ match idc in (ident t) return (Compile.value' true t) with
=>
match a with
| Some x' => UnderLets.Base (x' fv0)
- | None => default
+ | 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 => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_partition)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
end)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_partition)%expr @
+ (λ x1 : var (type.base A),
+ UnderLets.to_expr (x ($x1)))%expr @ x0)%expr_pat
end)
| @ident.List_fold_right A B =>
fun
@@ -1576,10 +1885,6 @@ 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))) =>
- llet default := 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 in
Compile.castv x
(UnderLets.UnderLets base.type ident var (defaults.expr (type.base A)))
(fun
@@ -1632,11 +1937,25 @@ match idc in (ident t) return (Compile.value' true t) with
(let (anyexpr_ty, _) := a0 in
anyexpr_ty))) := fv in
unwrap))
- | None => default
+ | 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 => default
+ | 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 => default
+ | 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)
| @ident.List_update_nth T =>
fun (x : defaults.expr (type.base base.type.nat))
@@ -1644,10 +1963,6 @@ 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.list T))) =>
- llet default := UnderLets.Base
- (#(ident.List_update_nth)%expr @ x @
- (λ x2 : var (type.base T),
- UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -1722,22 +2037,45 @@ match idc in (ident t) return (Compile.value' true t) with
a0 in
anyexpr_ty))) := fv in
unwrap))
- | None => default
+ | 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 => default
+ | 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 => default
+ | 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 => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_update_nth)%expr @ x @
+ (λ x2 : var (type.base T),
+ UnderLets.to_expr (x0 ($x2)))%expr @ x1)%expr_pat
end
- | _ => default
+ | ($_)%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)) =>
- llet default := UnderLets.Base
- (#(ident.List_nth_default)%expr @ x @ x0 @ x1)%expr_pat in
match x1 with
| #(idc)%expr_pat =>
match
@@ -1756,9 +2094,9 @@ match idc in (ident t) return (Compile.value' true t) with
Compile.castbe x
(UnderLets.UnderLets base.type ident var
(defaults.expr (type.base T)))
- (fun default0 : option (defaults.expr (type.base T)) =>
- match default0 with
- | Some default1 =>
+ (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
@@ -1775,20 +2113,30 @@ match idc in (ident t) return (Compile.value' true t) with
match a with
| Some x' =>
UnderLets.Base
- (x' (nth_default default1 ls0 args))
- | None => default
+ (x' (nth_default default0 ls0 args))
+ | None =>
+ UnderLets.Base
+ (#(ident.List_nth_default)%expr @ x @ x0 @
+ x1)%expr_pat
end)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_nth_default)%expr @ x @ x0 @ x1)%expr_pat
end)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_nth_default)%expr @ x @ x0 @ x1)%expr_pat
end)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.List_nth_default)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.List_nth_default)%expr @ x @ x0 @ x1)%expr_pat
end
| ident.Z_add =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x + x0)%expr in
match x with
| #(idc)%expr_pat =>
match
@@ -1821,17 +2169,16 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##((args + args0)%Z)%expr
- | None => default
+ | None => UnderLets.Base (x + x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x + x0)%expr
end
- | None => default
+ | None => UnderLets.Base (x + x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x + x0)%expr
end
| ident.Z_mul =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x * x0)%expr in
match x with
| #(idc)%expr_pat =>
match
@@ -1864,17 +2211,16 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##((args * args0)%Z)%expr
- | None => default
+ | None => UnderLets.Base (x * x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x * x0)%expr
end
- | None => default
+ | None => UnderLets.Base (x * x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x * x0)%expr
end
| ident.Z_pow =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -1907,17 +2253,17 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##(args ^ args0)%expr
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_pow)%expr @ x @ x0)%expr_pat
end
| ident.Z_sub =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x - x0)%expr in
match x with
| #(idc)%expr_pat =>
match
@@ -1950,17 +2296,16 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##((args - args0)%Z)%expr
- | None => default
+ | None => UnderLets.Base (x - x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x - x0)%expr
end
- | None => default
+ | None => UnderLets.Base (x - x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x - x0)%expr
end
| ident.Z_opp =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (- x)%expr in
match x with
| #(idc)%expr_pat =>
match
@@ -1976,13 +2321,12 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args => UnderLets.Base ##((- args)%Z)%expr
- | None => default
+ | None => UnderLets.Base (- x)%expr
end
- | _ => default
+ | _ => UnderLets.Base (- x)%expr
end
| ident.Z_div =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x / x0)%expr in
match x with
| #(idc)%expr_pat =>
match
@@ -2015,17 +2359,16 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##((args / args0)%Z)%expr
- | None => default
+ | None => UnderLets.Base (x / x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x / x0)%expr
end
- | None => default
+ | None => UnderLets.Base (x / x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x / x0)%expr
end
| ident.Z_modulo =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x mod x0)%expr in
match x with
| #(idc)%expr_pat =>
match
@@ -2058,17 +2401,16 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##((args mod args0)%Z)%expr
- | None => default
+ | None => UnderLets.Base (x mod x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x mod x0)%expr
end
- | None => default
+ | None => UnderLets.Base (x mod x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x mod x0)%expr
end
| ident.Z_log2 =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_log2)%expr @ x)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2084,13 +2426,12 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args => UnderLets.Base ##(Z.log2 args)%expr
- | None => default
+ | None => UnderLets.Base (#(ident.Z_log2)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_log2)%expr @ x)%expr_pat
end
| ident.Z_log2_up =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_log2_up)%expr @ x)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2106,13 +2447,12 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args => UnderLets.Base ##(Z.log2_up args)%expr
- | None => default
+ | None => UnderLets.Base (#(ident.Z_log2_up)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_log2_up)%expr @ x)%expr_pat
end
| ident.Z_eqb =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2145,17 +2485,17 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##(args =? args0)%expr
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_eqb)%expr @ x @ x0)%expr_pat
end
| ident.Z_leb =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2188,17 +2528,17 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##(args <=? args0)%expr
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_leb)%expr @ x @ x0)%expr_pat
end
| ident.Z_geb =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2231,17 +2571,17 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##(args >=? args0)%expr
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_geb)%expr @ x @ x0)%expr_pat
end
| ident.Z_of_nat =>
fun x : defaults.expr (type.base base.type.nat) =>
- llet default := UnderLets.Base (#(ident.Z_of_nat)%expr @ x)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2257,13 +2597,12 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args => UnderLets.Base ##(Z.of_nat args)%expr
- | None => default
+ | None => UnderLets.Base (#(ident.Z_of_nat)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_of_nat)%expr @ x)%expr_pat
end
| ident.Z_to_nat =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_to_nat)%expr @ x)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2279,13 +2618,12 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args => UnderLets.Base ##(Z.to_nat args)%expr
- | None => default
+ | None => UnderLets.Base (#(ident.Z_to_nat)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_to_nat)%expr @ x)%expr_pat
end
| ident.Z_shiftr =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x >> x0)%expr in
match x with
| #(idc)%expr_pat =>
match
@@ -2318,17 +2656,16 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##(Z.shiftr args args0)%expr
- | None => default
+ | None => UnderLets.Base (x >> x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x >> x0)%expr
end
- | None => default
+ | None => UnderLets.Base (x >> x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x >> x0)%expr
end
| ident.Z_shiftl =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x << x0)%expr in
match x with
| #(idc)%expr_pat =>
match
@@ -2361,17 +2698,16 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##(Z.shiftl args args0)%expr
- | None => default
+ | None => UnderLets.Base (x << x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x << x0)%expr
end
- | None => default
+ | None => UnderLets.Base (x << x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x << x0)%expr
end
| ident.Z_land =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x &' x0)%expr in
match x with
| #(idc)%expr_pat =>
match
@@ -2404,17 +2740,16 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##(Z.land args args0)%expr
- | None => default
+ | None => UnderLets.Base (x &' x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x &' x0)%expr
end
- | None => default
+ | None => UnderLets.Base (x &' x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x &' x0)%expr
end
| ident.Z_lor =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (x || x0)%expr in
match x with
| #(idc)%expr_pat =>
match
@@ -2447,17 +2782,16 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args0 => UnderLets.Base ##(Z.lor args args0)%expr
- | None => default
+ | None => UnderLets.Base (x || x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x || x0)%expr
end
- | None => default
+ | None => UnderLets.Base (x || x0)%expr
end
- | _ => default
+ | _ => UnderLets.Base (x || x0)%expr
end
| ident.Z_bneg =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_bneg)%expr @ x)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2473,14 +2807,12 @@ match idc in (ident t) return (Compile.value' true t) with
end
with
| Some args => UnderLets.Base ##(Definitions.Z.bneg args)%expr
- | None => default
+ | None => UnderLets.Base (#(ident.Z_bneg)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_bneg)%expr @ x)%expr_pat
end
| ident.Z_lnot_modulo =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_lnot_modulo)%expr @ x @ x0)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2515,18 +2847,21 @@ match idc in (ident t) return (Compile.value' true t) with
| Some args0 =>
UnderLets.Base
##(Definitions.Z.lnot_modulo args args0)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_lnot_modulo)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_lnot_modulo)%expr @ x @ x0)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Z_lnot_modulo)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_lnot_modulo)%expr @ x @ x0)%expr_pat
end
| ident.Z_mul_split =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2582,22 +2917,29 @@ match idc in (ident t) return (Compile.value' true t) with
'(a, b) :=
Definitions.Z.mul_split args args0 args1 in
(##(a)%expr, ##(b)%expr)%expr_pat)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
end
| ident.Z_add_get_carry =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2653,22 +2995,31 @@ match idc in (ident t) return (Compile.value' true t) with
'(a, b) :=
Definitions.Z.add_get_carry_full args args0
args1 in (##(a)%expr, ##(b)%expr)%expr_pat)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base (#(ident.Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
end
| ident.Z_add_with_carry =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2722,22 +3073,32 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base
##(Definitions.Z.add_with_carry args args0
args1)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
end
| ident.Z_add_with_get_carry =>
fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2815,26 +3176,44 @@ match idc in (ident t) return (Compile.value' true t) with
Definitions.Z.add_with_get_carry_full
args args0 args1 args2 in
(##(a)%expr, ##(b)%expr)%expr_pat)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @
+ x0 @ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @
+ x0 @ x1 @ x2)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @
+ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @
+ x2)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
| ident.Z_sub_get_borrow =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2890,22 +3269,32 @@ match idc in (ident t) return (Compile.value' true t) with
'(a, b) :=
Definitions.Z.sub_get_borrow_full args args0
args1 in (##(a)%expr, ##(b)%expr)%expr_pat)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
end
| ident.Z_sub_with_get_borrow =>
fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -2983,26 +3372,44 @@ match idc in (ident t) return (Compile.value' true t) with
Definitions.Z.sub_with_get_borrow_full
args args0 args1 args2 in
(##(a)%expr, ##(b)%expr)%expr_pat)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @
+ x @ x0 @ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @
+ x0 @ x1 @ x2)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @
+ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @
+ x2)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
| ident.Z_zselect =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -3055,22 +3462,29 @@ match idc in (ident t) return (Compile.value' true t) with
| Some args1 =>
UnderLets.Base
##(Definitions.Z.zselect args args0 args1)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
end
| ident.Z_add_modulo =>
fun x x0 x1 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -3123,22 +3537,30 @@ match idc in (ident t) return (Compile.value' true t) with
| Some args1 =>
UnderLets.Base
##(Definitions.Z.add_modulo args args0 args1)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
end
| ident.Z_rshi =>
fun x x0 x1 x2 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base
- (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -3213,25 +3635,38 @@ match idc in (ident t) return (Compile.value' true t) with
UnderLets.Base
##(Definitions.Z.rshi args args0 args1
args2)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @
+ x2)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
end
| ident.Z_cc_m =>
fun x x0 : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -3265,17 +3700,17 @@ match idc in (ident t) return (Compile.value' true t) with
with
| Some args0 =>
UnderLets.Base ##(Definitions.Z.cc_m args args0)%expr
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_cc_m)%expr @ x @ x0)%expr_pat
end
| ident.Z_cast range =>
fun x : defaults.expr (type.base base.type.Z) =>
- llet default := UnderLets.Base (#(ident.Z_cast range)%expr @ x)%expr_pat in
match x with
| #(idc)%expr_pat =>
match
@@ -3293,13 +3728,12 @@ match idc in (ident t) return (Compile.value' true t) with
| Some args =>
UnderLets.Base
##(ident.cast ident.cast_outside_of_range range args)%expr
- | None => default
+ | None => UnderLets.Base (#(ident.Z_cast range)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_cast range)%expr @ x)%expr_pat
end
| ident.Z_cast2 range =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat in
match x with
| (#(idc) @ x1 @ x0)%expr_pat =>
match
@@ -3354,22 +3788,25 @@ match idc in (ident t) return (Compile.value' true t) with
ident.cast ident.cast_outside_of_range r2 x3))
(args0, args1) in
(##(a)%expr, ##(b)%expr)%expr_pat)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.Z_cast2 range)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.Z_cast2 range)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.Z_cast2 range)%expr @ x)%expr_pat
end
| ident.fancy_add log2wordmax imm =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat in
match x with
| (#(idc) @ x1 @ x0)%expr_pat =>
match
@@ -3423,24 +3860,33 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.fancy_add log2wordmax imm)))
(args0, args1) in
(##(a)%expr, ##(b)%expr)%expr_pat)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base (#(ident.fancy_add log2wordmax imm)%expr @ x)%expr_pat
end
| ident.fancy_addc log2wordmax imm =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat in
match x with
| (#(idc) @ x1 @ x0)%expr_pat =>
match
@@ -3532,30 +3978,52 @@ match idc in (ident t) return (Compile.value' true t) with
log2wordmax imm)))
(args1, args2, args3) in
(##(a)%expr, ##(b)%expr)%expr_pat)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_addc log2wordmax
+ imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_addc log2wordmax imm)%expr @
+ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_addc log2wordmax imm)%expr @
+ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_addc log2wordmax imm)%expr @
+ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_addc log2wordmax imm)%expr @ x)%expr_pat
end
| ident.fancy_sub log2wordmax imm =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat in
match x with
| (#(idc) @ x1 @ x0)%expr_pat =>
match
@@ -3609,24 +4077,33 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.fancy_sub log2wordmax imm)))
(args0, args1) in
(##(a)%expr, ##(b)%expr)%expr_pat)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base (#(ident.fancy_sub log2wordmax imm)%expr @ x)%expr_pat
end
| ident.fancy_subb log2wordmax imm =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat in
match x with
| (#(idc) @ x1 @ x0)%expr_pat =>
match
@@ -3718,30 +4195,52 @@ match idc in (ident t) return (Compile.value' true t) with
log2wordmax imm)))
(args1, args2, args3) in
(##(a)%expr, ##(b)%expr)%expr_pat)
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_subb log2wordmax
+ imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_subb log2wordmax imm)%expr @
+ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_subb log2wordmax imm)%expr @
+ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_subb log2wordmax imm)%expr @
+ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_subb log2wordmax imm)%expr @ x)%expr_pat
end
| ident.fancy_mulll log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat in
match x with
| (#(idc) @ x1 @ x0)%expr_pat =>
match
@@ -3792,22 +4291,31 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.to_fancy
(ident.fancy_mulll log2wordmax)))
(args0, args1)%core)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base (#(ident.fancy_mulll log2wordmax)%expr @ x)%expr_pat
end
| ident.fancy_mullh log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat in
match x with
| (#(idc) @ x1 @ x0)%expr_pat =>
match
@@ -3858,22 +4366,31 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.to_fancy
(ident.fancy_mullh log2wordmax)))
(args0, args1)%core)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base (#(ident.fancy_mullh log2wordmax)%expr @ x)%expr_pat
end
| ident.fancy_mulhl log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat in
match x with
| (#(idc) @ x1 @ x0)%expr_pat =>
match
@@ -3924,22 +4441,31 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.to_fancy
(ident.fancy_mulhl log2wordmax)))
(args0, args1)%core)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base (#(ident.fancy_mulhl log2wordmax)%expr @ x)%expr_pat
end
| ident.fancy_mulhh log2wordmax =>
fun x : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat in
match x with
| (#(idc) @ x1 @ x0)%expr_pat =>
match
@@ -3990,22 +4516,31 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.to_fancy
(ident.fancy_mulhh log2wordmax)))
(args0, args1)%core)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base (#(ident.fancy_mulhh log2wordmax)%expr @ x)%expr_pat
end
| ident.fancy_rshi log2wordmax x =>
fun x0 : defaults.expr (type.base (base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat in
match x0 with
| (#(idc) @ x2 @ x1)%expr_pat =>
match
@@ -4056,23 +4591,33 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.to_fancy
(ident.fancy_rshi log2wordmax x)))
(args0, args1)%core)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base (#(ident.fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
end
| ident.fancy_selc =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat in
match x with
| (#(idc) @ x1 @ x0)%expr_pat =>
match
@@ -4160,32 +4705,43 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.to_fancy
ident.fancy_selc))
(args1, args2, args3)%core)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_selc)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_selc)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_selc)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_selc)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_selc)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_selc)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.fancy_selc)%expr @ x)%expr_pat
end
| ident.fancy_selm log2wordmax =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base
- (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat in
match x with
| (#(idc) @ x1 @ x0)%expr_pat =>
match
@@ -4274,31 +4830,51 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.fancy_selm
log2wordmax)))
(args1, args2, args3)%core)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_selm log2wordmax)%expr @
+ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_selm log2wordmax)%expr @
+ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_selm log2wordmax)%expr @
+ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.fancy_selm log2wordmax)%expr @ x)%expr_pat
end
| ident.fancy_sell =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat in
match x with
| (#(idc) @ x1 @ x0)%expr_pat =>
match
@@ -4386,31 +4962,43 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.to_fancy
ident.fancy_sell))
(args1, args2, args3)%core)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_sell)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_sell)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_sell)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_sell)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_sell)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_sell)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.fancy_sell)%expr @ x)%expr_pat
end
| ident.fancy_addm =>
fun
x : defaults.expr
(type.base (base.type.Z * base.type.Z * base.type.Z)%etype) =>
- llet default := UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat in
match x with
| (#(idc) @ x1 @ x0)%expr_pat =>
match
@@ -4498,25 +5086,38 @@ match idc in (ident t) return (Compile.value' true t) with
(ident.to_fancy
ident.fancy_addm))
(args1, args2, args3)%core)%expr
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_addm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_addm)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_addm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_addm)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base
+ (#(ident.fancy_addm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ =>
+ UnderLets.Base
+ (#(ident.fancy_addm)%expr @ x)%expr_pat
end
- | None => default
+ | None =>
+ UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat
end
- | None => default
+ | None => UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat
end
- | _ => default
+ | _ => UnderLets.Base (#(ident.fancy_addm)%expr @ x)%expr_pat
end
end
: Compile.value' true t