aboutsummaryrefslogtreecommitdiff
path: root/src
diff options
context:
space:
mode:
authorGravatar Jason Gross <jgross@mit.edu>2018-10-28 22:02:43 -0400
committerGravatar Jason Gross <jgross@mit.edu>2018-10-28 22:02:43 -0400
commit841b485536d21947debb9797205f7d435aaff980 (patch)
treec53d95148c5f24b92eb1fc2e6b48d857dbf039fb /src
parentd5298f223b503d8ca02c6869ba269bf8f4b84b96 (diff)
Add an extra check in the rewriter
Diffstat (limited to 'src')
-rw-r--r--src/Experiments/NewPipeline/Rewriter.v27
-rw-r--r--src/Experiments/NewPipeline/arith_rewrite_head.out1920
-rw-r--r--src/Experiments/NewPipeline/fancy_rewrite_head.out1694
-rw-r--r--src/Experiments/NewPipeline/nbe_rewrite_head.out3759
4 files changed, 4288 insertions, 3112 deletions
diff --git a/src/Experiments/NewPipeline/Rewriter.v b/src/Experiments/NewPipeline/Rewriter.v
index 4882b9b1e..484b8ab05 100644
--- a/src/Experiments/NewPipeline/Rewriter.v
+++ b/src/Experiments/NewPipeline/Rewriter.v
@@ -78,7 +78,8 @@ Module Compilers.
Fixpoint add_var_types_cps {t : type} (v : var_types_of t) (evm : EvarMap) : ~> EvarMap
:= fun T k
=> match t return var_types_of t -> T with
- | type.var p => fun t => k (PositiveMap.add p t evm)
+ | type.var p
+ => fun t => k (PositiveMap.add p t evm)
| type.prod A B
=> fun '(a, b) => @add_var_types_cps A a evm _ (fun evm => @add_var_types_cps B b evm _ k)
| type.list A => fun t => @add_var_types_cps A t evm _ k
@@ -749,12 +750,22 @@ Module Compilers.
=> None
end%option.
+ (* for unfolding help *)
+ Definition type_type_beq := (type.type_beq _ base.type.type_beq).
+
Definition unify_types {t} (e : rawexpr) (p : pattern t) : ~> option EvarMap
:= fun T k
=> match preunify_types e p with
| Some (Some (pt, t))
=> match pattern.type.unify_extracted pt t with
- | Some vars => pattern.type.add_var_types_cps vars (PositiveMap.empty _) _ (fun evm => k (Some evm))
+ | Some vars
+ => pattern.type.add_var_types_cps
+ vars (PositiveMap.empty _) _
+ (fun evm
+ => (* there might be multiple type variables which map to incompatible types; we check for that here *)
+ if type_type_beq (pattern.type.subst_default pt evm) t
+ then k (Some evm)
+ else k None)
| None => k None
end
| Some None
@@ -2197,6 +2208,7 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
base.interp base.try_make_transport_cps
type.try_make_transport_cps type.try_transport_cps
pattern.type.unify_extracted_cps
+ Compile.type_type_beq
Let_In Option.sequence Option.sequence_return
UnderLets.splice UnderLets.to_expr
Compile.option_bind' pident_unify_unknown invert_bind_args_unknown Compile.normalize_deep_rewrite_rule
@@ -2230,6 +2242,7 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
Compile.rValueOrExpr
Compile.swap_list
Compile.type_of_rawexpr
+ Compile.type_type_beq
Compile.value
(*Compile.value'*)
Compile.value_of_rawexpr
@@ -2308,7 +2321,7 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
Local Set Printing Depth 1000000.
Local Set Printing Width 200.
Import RewriterPrintingNotations.
- (* Redirect "/tmp/fancy_rewrite_head" Print fancy_rewrite_head. *)
+ Redirect "fancy_rewrite_head" Print fancy_rewrite_head.
End red_fancy.
Section red_fancy_with_casts.
Context (invert_low invert_high : Z (*log2wordmax*) -> Z -> @option Z)
@@ -2322,7 +2335,7 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
Local Set Printing Depth 1000000.
Local Set Printing Width 200.
Import RewriterPrintingNotations.
- (* Redirect "/tmp/fancy_with_casts_rewrite_head" Print fancy_with_casts_rewrite_head. *)
+ Redirect "fancy_with_casts_rewrite_head" Print fancy_with_casts_rewrite_head.
End red_fancy_with_casts.
Section red_nbe.
Context {var : type.type base.type -> Type}
@@ -2339,7 +2352,7 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
Local Set Printing Depth 1000000.
Local Set Printing Width 200.
Import RewriterPrintingNotations.
- (* Redirect "/tmp/nbe_rewrite_head" Print nbe_rewrite_head. *)
+ Redirect "nbe_rewrite_head" Print nbe_rewrite_head.
End red_nbe.
Section red_arith.
@@ -2359,7 +2372,7 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
Local Set Printing Depth 1000000.
Local Set Printing Width 200.
Import RewriterPrintingNotations.
- (* Redirect "/tmp/arith_rewrite_head" Print arith_rewrite_head. *)
+ Redirect "arith_rewrite_head" Print arith_rewrite_head.
End red_arith.
Section red_arith_with_casts.
@@ -2374,7 +2387,7 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
Local Set Printing Depth 1000000.
Local Set Printing Width 200.
Import RewriterPrintingNotations.
- (* Redirect "/tmp/arith_with_casts_rewrite_head" Print arith_with_casts_rewrite_head. *)
+ Redirect "arith_with_casts_rewrite_head" Print arith_with_casts_rewrite_head.
End red_arith_with_casts.
Definition RewriteNBE {t} (e : expr.Expr (ident:=ident) t) : expr.Expr (ident:=ident) t
diff --git a/src/Experiments/NewPipeline/arith_rewrite_head.out b/src/Experiments/NewPipeline/arith_rewrite_head.out
index fe90d714c..112dd2250 100644
--- a/src/Experiments/NewPipeline/arith_rewrite_head.out
+++ b/src/Experiments/NewPipeline/arith_rewrite_head.out
@@ -28,14 +28,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
s0) -> s)%ptype option (fun x2 : option => x2)
with
| Some (_, _, _, (_, (_, (_, _)), b3, b2)) =>
- _ <- ident.unify pattern.ident.fst fst;
- _ <- ident.unify pattern.ident.pair pair;
- v <- type.try_make_transport_cps s0 b3;
- _ <- type.try_make_transport_cps s b2;
- v1 <- base.try_make_transport_cps b3 A;
- v2 <- base.try_make_transport_cps A A;
- v3 <- base.try_make_transport_cps A A;
- Some (Base (v3 (v2 (v1 (v (Compile.reflect x1))))))
+ if
+ type.type_beq base.type base.type.type_beq
+ (((b3 * b2)%etype -> b3) ->
+ ((b3 -> b2 -> (b3 * b2)%etype) -> b3) -> b2)%ptype
+ (((A * B)%etype -> A) ->
+ (((let (x2, _) := args in x2) ->
+ (let (_, y) := args in y) ->
+ ((let (x2, _) := args in x2) * (let (_, y) := args in y))%etype) ->
+ s0) -> s)%ptype
+ then
+ _ <- ident.unify pattern.ident.fst fst;
+ _ <- ident.unify pattern.ident.pair pair;
+ v <- type.try_make_transport_cps s0 b3;
+ _ <- type.try_make_transport_cps s b2;
+ v1 <- base.try_make_transport_cps b3 A;
+ v2 <- base.try_make_transport_cps A A;
+ v3 <- base.try_make_transport_cps A A;
+ Some (Base (v3 (v2 (v1 (v (Compile.reflect x1))))))
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
@@ -68,14 +79,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
s0) -> s)%ptype option (fun x2 : option => x2)
with
| Some (_, _, _, (_, (_, (_, _)), b3, b2)) =>
- _ <- ident.unify pattern.ident.snd snd;
- _ <- ident.unify pattern.ident.pair pair;
- _ <- type.try_make_transport_cps s0 b3;
- v0 <- type.try_make_transport_cps s b2;
- v1 <- base.try_make_transport_cps b2 B;
- v2 <- base.try_make_transport_cps B B;
- v3 <- base.try_make_transport_cps B B;
- Some (Base (v3 (v2 (v1 (v0 (Compile.reflect x0))))))
+ if
+ type.type_beq base.type base.type.type_beq
+ (((b3 * b2)%etype -> b2) ->
+ ((b3 -> b2 -> (b3 * b2)%etype) -> b3) -> b2)%ptype
+ (((A * B)%etype -> B) ->
+ (((let (x2, _) := args in x2) ->
+ (let (_, y) := args in y) ->
+ ((let (x2, _) := args in x2) * (let (_, y) := args in y))%etype) ->
+ s0) -> s)%ptype
+ then
+ _ <- ident.unify pattern.ident.snd snd;
+ _ <- ident.unify pattern.ident.pair pair;
+ _ <- type.try_make_transport_cps s0 b3;
+ v0 <- type.try_make_transport_cps s b2;
+ v1 <- base.try_make_transport_cps b2 B;
+ v2 <- base.try_make_transport_cps B B;
+ v3 <- base.try_make_transport_cps B B;
+ Some (Base (v3 (v2 (v1 (v0 (Compile.reflect x0))))))
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
@@ -201,11 +223,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> ℤ)%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? 0
- then Some x0
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? 0
+ then Some x0
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| _ => None
@@ -218,11 +245,17 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ℤ -> (projT1 args))%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? 0
- then Some x
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? 0
+ then Some x
+ else None);
+ Some (Base x1)
+ else None
| None => None
end);;
match x with
@@ -235,16 +268,21 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- x2 <- (if (let (x2, _) := idc_args in x2) >? 0
- then
- Some
- (##(let (x2, _) := idc_args in x2) -
- v (Compile.reflect x1))%expr
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ)%ptype (s -> (projT1 args0))%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ x2 <- (if (let (x2, _) := idc_args in x2) >? 0
+ then
+ Some
+ (##(let (x2, _) := idc_args in x2) -
+ v (Compile.reflect x1))%expr
+ else None);
+ Some (Base x2)
+ else None
| None => None
end);;
_ <- invert_bind_args idc0 Raw.ident.Z_opp;
@@ -254,17 +292,22 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(s -> (projT1 args0))%ptype option (fun x2 : option => x2)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- x2 <- (if (let (x2, _) := idc_args in x2) <? 0
- then
- Some
- (-
- (v (Compile.reflect x1) +
- ##(- (let (x2, _) := idc_args in x2))%Z))%expr
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (s -> (projT1 args0))%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ x2 <- (if (let (x2, _) := idc_args in x2) <? 0
+ then
+ Some
+ (-
+ (v (Compile.reflect x1) +
+ ##(- (let (x2, _) := idc_args in x2))%Z))%expr
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -282,16 +325,21 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> s)%ptype option (fun x2 : option => x2)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v <- type.try_make_transport_cps s ℤ;
- x2 <- (if (let (x2, _) := idc_args in x2) >? 0
- then
- Some
- (##(let (x2, _) := idc_args in x2) -
- v (Compile.reflect x1))%expr
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ)%ptype ((projT1 args) -> s)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v <- type.try_make_transport_cps s ℤ;
+ x2 <- (if (let (x2, _) := idc_args in x2) >? 0
+ then
+ Some
+ (##(let (x2, _) := idc_args in x2) -
+ v (Compile.reflect x1))%expr
+ else None);
+ Some (Base x2)
+ else None
| None => None
end);;
args <- invert_bind_args idc0 Raw.ident.Literal;
@@ -301,17 +349,22 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> s)%ptype option (fun x2 : option => x2)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v <- type.try_make_transport_cps s ℤ;
- x2 <- (if (let (x2, _) := idc_args in x2) <? 0
- then
- Some
- (-
- (##(- (let (x2, _) := idc_args in x2))%Z +
- v (Compile.reflect x1)))%expr
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> s)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v <- type.try_make_transport_cps s ℤ;
+ x2 <- (if (let (x2, _) := idc_args in x2) <? 0
+ then
+ Some
+ (-
+ (##(- (let (x2, _) := idc_args in x2))%Z +
+ v (Compile.reflect x1)))%expr
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| @expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x2 =>
@@ -322,11 +375,17 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(s0 -> s)%ptype option (fun x3 : option => x3)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s0 ℤ;
- v0 <- type.try_make_transport_cps s ℤ;
- Some
- (Base
- (- (v (Compile.reflect x2) + v0 (Compile.reflect x1)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (s0 -> s)%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ v0 <- type.try_make_transport_cps s ℤ;
+ Some
+ (Base
+ (-
+ (v (Compile.reflect x2) + v0 (Compile.reflect x1)))%expr)
+ else None
| None => None
end
| @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
@@ -348,8 +407,13 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x2 : option => x2)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s ℤ;
- Some (Base (x0 - v (Compile.reflect x1))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (s -> ℤ)%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ Some (Base (x0 - v (Compile.reflect x1))%expr)
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -365,8 +429,13 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x2 : option => x2)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s ℤ;
- Some (Base (x - v (Compile.reflect x1))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> s)%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ Some (Base (x - v (Compile.reflect x1))%expr)
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -390,14 +459,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##((let (x1, _) := idc_args in x1) *
- (let (x1, _) := idc_args0 in x1))%Z)%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##((let (x1, _) := idc_args in x1) *
+ (let (x1, _) := idc_args0 in x1))%Z)%expr)
+ else None
| None => None
end
| _ => None
@@ -408,11 +482,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> ℤ)%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? 0
- then Some (##0)%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? 0
+ then Some (##0)%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| _ => None
@@ -425,11 +504,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ℤ -> (projT1 args))%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? 0
- then Some (##0)%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? 0
+ then Some (##0)%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| _ => None
@@ -442,11 +526,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> ℤ)%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? 1
- then Some x0
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? 1
+ then Some x0
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| _ => None
@@ -459,11 +548,17 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ℤ -> (projT1 args))%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? 1
- then Some x
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? 1
+ then Some x
+ else None);
+ Some (Base x1)
+ else None
| None => None
end);;
match x with
@@ -475,13 +570,18 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(s -> (projT1 args0))%ptype option (fun x2 : option => x2)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- x2 <- (if (let (x2, _) := idc_args in x2) =? -1
- then Some (v (Compile.reflect x1))
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (s -> (projT1 args0))%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ x2 <- (if (let (x2, _) := idc_args in x2) =? -1
+ then Some (v (Compile.reflect x1))
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -499,13 +599,18 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> s)%ptype option (fun x2 : option => x2)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v <- type.try_make_transport_cps s ℤ;
- x2 <- (if (let (x2, _) := idc_args in x2) =? -1
- then Some (v (Compile.reflect x1))
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> s)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v <- type.try_make_transport_cps s ℤ;
+ x2 <- (if (let (x2, _) := idc_args in x2) =? -1
+ then Some (v (Compile.reflect x1))
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -523,11 +628,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> ℤ)%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? -1
- then Some (- x0)%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? -1
+ then Some (- x0)%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| _ => None
@@ -540,11 +650,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ℤ -> (projT1 args))%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? -1
- then Some (- x)%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? -1
+ then Some (- x)%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| _ => None
@@ -557,13 +672,18 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> ℤ)%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) <? 0
- then
- Some
- (- (##(- (let (x1, _) := idc_args in x1))%Z * x0))%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) <? 0
+ then
+ Some
+ (- (##(- (let (x1, _) := idc_args in x1))%Z * x0))%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| _ => None
@@ -576,13 +696,18 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ℤ -> (projT1 args))%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) <? 0
- then
- Some
- (- (x * ##(- (let (x1, _) := idc_args in x1))%Z))%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) <? 0
+ then
+ Some
+ (- (x * ##(- (let (x1, _) := idc_args in x1))%Z))%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
@@ -595,11 +720,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(s0 -> s)%ptype option (fun x3 : option => x3)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s0 ℤ;
- v0 <- type.try_make_transport_cps s ℤ;
- Some
- (Base
- (v (Compile.reflect x2) * v0 (Compile.reflect x1))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (s0 -> s)%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ v0 <- type.try_make_transport_cps s ℤ;
+ Some
+ (Base
+ (v (Compile.reflect x2) * v0 (Compile.reflect x1))%expr)
+ else None
| None => None
end
| @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
@@ -621,8 +751,13 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x2 : option => x2)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s ℤ;
- Some (Base (- (v (Compile.reflect x1) * x0))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (s -> ℤ)%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ Some (Base (- (v (Compile.reflect x1) * x0))%expr)
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -638,16 +773,21 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ℤ -> (projT1 args))%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if
- ((let (x1, _) := idc_args in x1) =?
- 2 ^ Z.log2 (let (x1, _) := idc_args in x1)) &&
- negb ((let (x1, _) := idc_args in x1) =? 2)
- then
- Some
- (x << ##(Z.log2 (let (x1, _) := idc_args in x1)))%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if
+ ((let (x1, _) := idc_args in x1) =?
+ 2 ^ Z.log2 (let (x1, _) := idc_args in x1)) &&
+ negb ((let (x1, _) := idc_args in x1) =? 2)
+ then
+ Some
+ (x << ##(Z.log2 (let (x1, _) := idc_args in x1)))%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
@@ -657,8 +797,13 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x2 : option => x2)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s ℤ;
- Some (Base (- (x * v (Compile.reflect x1)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> s)%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ Some (Base (- (x * v (Compile.reflect x1)))%expr)
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -674,16 +819,22 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> ℤ)%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if
- ((let (x1, _) := idc_args in x1) =?
- 2 ^ Z.log2 (let (x1, _) := idc_args in x1)) &&
- negb ((let (x1, _) := idc_args in x1) =? 2)
- then
- Some
- (x0 << ##(Z.log2 (let (x1, _) := idc_args in x1)))%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x1 <- (if
+ ((let (x1, _) := idc_args in x1) =?
+ 2 ^ Z.log2 (let (x1, _) := idc_args in x1)) &&
+ negb ((let (x1, _) := idc_args in x1) =? 2)
+ then
+ Some
+ (x0 << ##(Z.log2 (let (x1, _) := idc_args in x1)))%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end);;
match x0 with
@@ -706,25 +857,31 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, (_, (_, _))) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s2 ℤ;
- v0 <- type.try_make_transport_cps s1 ℤ;
- x5 <- (if
- (Z.abs (let (x5, _) := idc_args in x5) <=?
- Z.abs max_const_val) &&
- (Z.abs (let (x5, _) := idc_args0 in x5) <=?
- Z.abs max_const_val)
- then
- Some
- (v (Compile.reflect x4) *
- (v0 (Compile.reflect x3) *
- (##(let (x5, _) := idc_args in x5) *
- ##(let (x5, _) := idc_args0 in x5))))%expr
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args2) -> (projT1 args0) -> s2 -> s1)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args2);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ v <- type.try_make_transport_cps s2 ℤ;
+ v0 <- type.try_make_transport_cps s1 ℤ;
+ x5 <- (if
+ (Z.abs (let (x5, _) := idc_args in x5) <=?
+ Z.abs max_const_val) &&
+ (Z.abs (let (x5, _) := idc_args0 in x5) <=?
+ Z.abs max_const_val)
+ then
+ Some
+ (v (Compile.reflect x4) *
+ (v0 (Compile.reflect x3) *
+ (##(let (x5, _) := idc_args in x5) *
+ ##(let (x5, _) := idc_args0 in x5))))%expr
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| @expr.App _ _ _ s1 _ (@expr.App _ _ _ s2 _ ($_)%expr _)
@@ -757,25 +914,31 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, (_, (_, _))) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s0 ℤ;
- v0 <- type.try_make_transport_cps s2 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x5 <- (if
- (Z.abs (let (x5, _) := idc_args in x5) <=?
- Z.abs max_const_val) &&
- (Z.abs (let (x5, _) := idc_args0 in x5) <=?
- Z.abs max_const_val)
- then
- Some
- (v (Compile.reflect x2) *
- (v0 (Compile.reflect x4) *
- (##(let (x5, _) := idc_args in x5) *
- ##(let (x5, _) := idc_args0 in x5))))%expr
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args2) -> s0 -> s2 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args2);
+ v <- type.try_make_transport_cps s0 ℤ;
+ v0 <- type.try_make_transport_cps s2 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x5 <- (if
+ (Z.abs (let (x5, _) := idc_args in x5) <=?
+ Z.abs max_const_val) &&
+ (Z.abs (let (x5, _) := idc_args0 in x5) <=?
+ Z.abs max_const_val)
+ then
+ Some
+ (v (Compile.reflect x2) *
+ (v0 (Compile.reflect x4) *
+ (##(let (x5, _) := idc_args in x5) *
+ ##(let (x5, _) := idc_args0 in x5))))%expr
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| (@expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t1 idc1) x4 @
@@ -801,20 +964,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s0 ℤ;
- v0 <- type.try_make_transport_cps s ℤ;
- x3 <- (if
- Z.abs (let (x3, _) := idc_args in x3) <=?
- Z.abs max_const_val
- then
- Some
- (v (Compile.reflect x2) *
- (v0 (Compile.reflect x1) *
- ##(let (x3, _) := idc_args in x3)))%expr
- else None);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ -> ℤ)%ptype ((projT1 args0) -> s0 -> s)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ v <- type.try_make_transport_cps s0 ℤ;
+ v0 <- type.try_make_transport_cps s ℤ;
+ x3 <- (if
+ Z.abs (let (x3, _) := idc_args in x3) <=?
+ Z.abs max_const_val
+ then
+ Some
+ (v (Compile.reflect x2) *
+ (v0 (Compile.reflect x1) *
+ ##(let (x3, _) := idc_args in x3)))%expr
+ else None);
+ Some (Base x3)
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ |
@@ -834,13 +1002,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> ℤ)%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if
- Z.abs (let (x1, _) := idc_args in x1) <=?
- Z.abs max_const_val
- then Some (x0 * ##(let (x1, _) := idc_args in x1))%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if
+ Z.abs (let (x1, _) := idc_args in x1) <=?
+ Z.abs max_const_val
+ then
+ Some (x0 * ##(let (x1, _) := idc_args in x1))%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| _ => None
@@ -861,13 +1035,18 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args0) -> s)%ptype option (fun x2 : option => x2)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s ℤ;
- x2 <- (if (let (x2, _) := idc_args in x2) =? 0
- then Some (v (Compile.reflect x1))
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> s)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ v <- type.try_make_transport_cps s ℤ;
+ x2 <- (if (let (x2, _) := idc_args in x2) =? 0
+ then Some (v (Compile.reflect x1))
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -881,11 +1060,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> ℤ)%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? 0
- then Some (- x0)%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? 0
+ then Some (- x0)%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| _ => None
@@ -898,11 +1082,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ℤ -> (projT1 args))%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? 0
- then Some x
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? 0
+ then Some x
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
@@ -915,16 +1104,21 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> s)%ptype option (fun x2 : option => x2)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v <- type.try_make_transport_cps s ℤ;
- x2 <- (if (let (x2, _) := idc_args in x2) >? 0
- then
- Some
- (##(let (x2, _) := idc_args in x2) +
- v (Compile.reflect x1))%expr
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ)%ptype ((projT1 args) -> s)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v <- type.try_make_transport_cps s ℤ;
+ x2 <- (if (let (x2, _) := idc_args in x2) >? 0
+ then
+ Some
+ (##(let (x2, _) := idc_args in x2) +
+ v (Compile.reflect x1))%expr
+ else None);
+ Some (Base x2)
+ else None
| None => None
end);;
args <- invert_bind_args idc0 Raw.ident.Literal;
@@ -934,16 +1128,21 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> s)%ptype option (fun x2 : option => x2)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v <- type.try_make_transport_cps s ℤ;
- x2 <- (if (let (x2, _) := idc_args in x2) <? 0
- then
- Some
- (v (Compile.reflect x1) -
- ##(- (let (x2, _) := idc_args in x2))%Z)%expr
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> s)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v <- type.try_make_transport_cps s ℤ;
+ x2 <- (if (let (x2, _) := idc_args in x2) <? 0
+ then
+ Some
+ (v (Compile.reflect x1) -
+ ##(- (let (x2, _) := idc_args in x2))%Z)%expr
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -961,13 +1160,18 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> ℤ)%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) <? 0
- then
- Some
- (- (##(- (let (x1, _) := idc_args in x1))%Z + x0))%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) <? 0
+ then
+ Some
+ (- (##(- (let (x1, _) := idc_args in x1))%Z + x0))%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
@@ -980,17 +1184,22 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(s -> (projT1 args))%ptype option (fun x2 : option => x2)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x2 <- (if (let (x2, _) := idc_args in x2) >? 0
- then
- Some
- (-
- (v (Compile.reflect x1) +
- ##(let (x2, _) := idc_args in x2)))%expr
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ)%ptype (s -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x2 <- (if (let (x2, _) := idc_args in x2) >? 0
+ then
+ Some
+ (-
+ (v (Compile.reflect x1) +
+ ##(let (x2, _) := idc_args in x2)))%expr
+ else None);
+ Some (Base x2)
+ else None
| None => None
end);;
args <- invert_bind_args idc0 Raw.ident.Literal;
@@ -1000,16 +1209,21 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(s -> (projT1 args))%ptype option (fun x2 : option => x2)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x2 <- (if (let (x2, _) := idc_args in x2) <? 0
- then
- Some
- (##(- (let (x2, _) := idc_args in x2))%Z -
- v (Compile.reflect x1))%expr
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (s -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x2 <- (if (let (x2, _) := idc_args in x2) <? 0
+ then
+ Some
+ (##(- (let (x2, _) := idc_args in x2))%Z -
+ v (Compile.reflect x1))%expr
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -1027,13 +1241,18 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ℤ -> (projT1 args))%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) <? 0
- then
- Some
- (x + ##(- (let (x1, _) := idc_args in x1))%Z)%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) <? 0
+ then
+ Some
+ (x + ##(- (let (x1, _) := idc_args in x1))%Z)%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x1 =>
@@ -1046,11 +1265,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(s0 -> s)%ptype option (fun x3 : option => x3)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s0 ℤ;
- v0 <- type.try_make_transport_cps s ℤ;
- Some
- (Base
- (v0 (Compile.reflect x1) - v (Compile.reflect x2))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (s0 -> s)%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ v0 <- type.try_make_transport_cps s ℤ;
+ Some
+ (Base
+ (v0 (Compile.reflect x1) - v (Compile.reflect x2))%expr)
+ else None
| None => None
end
| @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
@@ -1072,8 +1296,13 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x2 : option => x2)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s ℤ;
- Some (Base (- (v (Compile.reflect x1) + x0))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (s -> ℤ)%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ Some (Base (- (v (Compile.reflect x1) + x0))%expr)
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -1089,8 +1318,13 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x2 : option => x2)
with
| Some (_, _) =>
- v <- type.try_make_transport_cps s ℤ;
- Some (Base (x + v (Compile.reflect x1))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> s)%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ Some (Base (x + v (Compile.reflect x1))%expr)
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -1110,8 +1344,11 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some _ =>
- v <- type.try_make_transport_cps s ℤ;
- Some (Base (v (Compile.reflect x0)))
+ if type.type_beq base.type base.type.type_beq ℤ s
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ Some (Base (v (Compile.reflect x0)))
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -1123,18 +1360,21 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
pattern.type.unify_extracted_cps ℤ ℤ option (fun x0 : option => x0)
with
| Some _ =>
- fv <- (if negb (SubstVarLike.is_var_fst_snd_pair_opp x)
- then
- Some
- (UnderLet x
- (fun
- v : var
- (base.subst_default ℤ
- (PositiveMap.empty base.type)) =>
- Base (- $v)%expr))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
+ if type.type_beq base.type base.type.type_beq ℤ ℤ
+ then
+ fv <- (if negb (SubstVarLike.is_var_fst_snd_pair_opp x)
+ then
+ Some
+ (UnderLet x
+ (fun
+ v : var
+ (base.subst_default ℤ
+ (PositiveMap.empty base.type)) =>
+ Base (- $v)%expr))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
| None => None
end);;
None);;;
@@ -1149,11 +1389,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ℤ -> (projT1 args))%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? 1
- then Some x
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? 1
+ then Some x
+ else None);
+ Some (Base x1)
+ else None
| None => None
end);;
args <- invert_bind_args idc Raw.ident.Literal;
@@ -1162,15 +1407,20 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ℤ -> (projT1 args))%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if
- (let (x1, _) := idc_args in x1) =?
- 2 ^ Z.log2 (let (x1, _) := idc_args in x1)
- then
- Some
- (x >> ##(Z.log2 (let (x1, _) := idc_args in x1)))%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if
+ (let (x1, _) := idc_args in x1) =?
+ 2 ^ Z.log2 (let (x1, _) := idc_args in x1)
+ then
+ Some
+ (x >> ##(Z.log2 (let (x1, _) := idc_args in x1)))%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| _ => None
@@ -1187,11 +1437,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ℤ -> (projT1 args))%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? 1
- then Some (##0)%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? 1
+ then Some (##0)%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end);;
args <- invert_bind_args idc Raw.ident.Literal;
@@ -1200,15 +1455,20 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ℤ -> (projT1 args))%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if
- (let (x1, _) := idc_args in x1) =?
- 2 ^ Z.log2 (let (x1, _) := idc_args in x1)
- then
- Some
- (x &' ##((let (x1, _) := idc_args in x1) - 1)%Z)%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if
+ (let (x1, _) := idc_args in x1) =?
+ 2 ^ Z.log2 (let (x1, _) := idc_args in x1)
+ then
+ Some
+ (x &' ##((let (x1, _) := idc_args in x1) - 1)%Z)%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| _ => None
@@ -1233,11 +1493,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ℤ -> (projT1 args))%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? 0
- then Some (##0)%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? 0
+ then Some (##0)%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| _ => None
@@ -1250,11 +1515,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> ℤ)%ptype option (fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x1 <- (if (let (x1, _) := idc_args in x1) =? 0
- then Some (##0)%expr
- else None);
- Some (Base x1)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x1 <- (if (let (x1, _) := idc_args in x1) =? 0
+ then Some (##0)%expr
+ else None);
+ Some (Base x1)
+ else None
| None => None
end
| _ => None
@@ -1279,13 +1549,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, _) =>
- _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x2 <- (if (let (x2, _) := idc_args0 in x2) =? 0
- then Some ((##0)%expr, (##0)%expr)%expr_pat
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args0) -> (projT1 args)) -> ℤ)%ptype
+ then
+ _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x2 <- (if (let (x2, _) := idc_args0 in x2) =? 0
+ then Some ((##0)%expr, (##0)%expr)%expr_pat
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -1300,13 +1576,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, _) =>
- _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x2 <- (if (let (x2, _) := idc_args0 in x2) =? 0
- then Some ((##0)%expr, (##0)%expr)%expr_pat
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args0) -> ℤ) -> (projT1 args))%ptype
+ then
+ _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x2 <- (if (let (x2, _) := idc_args0 in x2) =? 0
+ then Some ((##0)%expr, (##0)%expr)%expr_pat
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -1321,13 +1603,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, _) =>
- _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x2 <- (if (let (x2, _) := idc_args0 in x2) =? 1
- then Some (x1, (##0)%expr)%expr_pat
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args0) -> (projT1 args)) -> ℤ)%ptype
+ then
+ _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x2 <- (if (let (x2, _) := idc_args0 in x2) =? 1
+ then Some (x1, (##0)%expr)%expr_pat
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -1342,13 +1630,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, _) =>
- _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x2 <- (if (let (x2, _) := idc_args0 in x2) =? 1
- then Some (x0, (##0)%expr)%expr_pat
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args0) -> ℤ) -> (projT1 args))%ptype
+ then
+ _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x2 <- (if (let (x2, _) := idc_args0 in x2) =? 1
+ then Some (x0, (##0)%expr)%expr_pat
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -1363,13 +1657,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, _) =>
- _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x2 <- (if (let (x2, _) := idc_args0 in x2) =? -1
- then Some ((- x1)%expr, (##0)%expr)%expr_pat
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args0) -> (projT1 args)) -> ℤ)%ptype
+ then
+ _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x2 <- (if (let (x2, _) := idc_args0 in x2) =? -1
+ then Some ((- x1)%expr, (##0)%expr)%expr_pat
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -1384,13 +1684,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, _) =>
- _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x2 <- (if (let (x2, _) := idc_args0 in x2) =? -1
- then Some ((- x0)%expr, (##0)%expr)%expr_pat
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args0) -> ℤ) -> (projT1 args))%ptype
+ then
+ _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x2 <- (if (let (x2, _) := idc_args0 in x2) =? -1
+ then Some ((- x0)%expr, (##0)%expr)%expr_pat
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -1402,11 +1708,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((ℤ -> ℤ) -> ℤ)%ptype option (fun x2 : option => x2)
with
| Some (_, _, _) =>
- Some
- (UnderLet (#(Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($v)%expr, #(snd)%expr @ ($v)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq ((ℤ -> ℤ) -> ℤ)%ptype
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ then
+ Some
+ (UnderLet (#(Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
+ (fun v : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($v)%expr, #(snd)%expr @ ($v)%expr)%expr_pat))
+ else None
| None => None
end);;
None);;;
@@ -1421,15 +1732,20 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((ℤ -> s) -> ℤ)%ptype option (fun x3 : option => x3)
with
| Some (_, _, _) =>
- v <- type.try_make_transport_cps s ℤ;
- Some
- (UnderLet
- (#(Z_sub_get_borrow)%expr @ x @ x1 @
- v (Compile.reflect x2))%expr_pat
- (fun v0 : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($v0)%expr,
- (- (#(snd)%expr @ $v0)%expr_pat)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype ((ℤ -> s) -> ℤ)%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ Some
+ (UnderLet
+ (#(Z_sub_get_borrow)%expr @ x @ x1 @
+ v (Compile.reflect x2))%expr_pat
+ (fun v0 : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($v0)%expr,
+ (- (#(snd)%expr @ $v0)%expr_pat)%expr)%expr_pat))
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -1445,15 +1761,20 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((ℤ -> ℤ) -> s)%ptype option (fun x3 : option => x3)
with
| Some (_, _, _) =>
- v <- type.try_make_transport_cps s ℤ;
- Some
- (UnderLet
- (#(Z_sub_get_borrow)%expr @ x @ x0 @
- v (Compile.reflect x2))%expr_pat
- (fun v0 : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($v0)%expr,
- (- (#(snd)%expr @ $v0)%expr_pat)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype ((ℤ -> ℤ) -> s)%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ Some
+ (UnderLet
+ (#(Z_sub_get_borrow)%expr @ x @ x0 @
+ v (Compile.reflect x2))%expr_pat
+ (fun v0 : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($v0)%expr,
+ (- (#(snd)%expr @ $v0)%expr_pat)%expr)%expr_pat))
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -1470,20 +1791,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (if (let (x2, _) := idc_args in x2) <? 0
- then
- Some
- (UnderLet
- (#(Z_sub_get_borrow)%expr @ x @ x1 @
- (##(- (let (x2, _) := idc_args in x2))%Z)%expr)%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($vc)%expr,
- (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype ((ℤ -> (projT1 args)) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ fv <- (if (let (x2, _) := idc_args in x2) <? 0
+ then
+ Some
+ (UnderLet
+ (#(Z_sub_get_borrow)%expr @ x @ x1 @
+ (##(- (let (x2, _) := idc_args in x2))%Z)%expr)%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($vc)%expr,
+ (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
| None => None
end
| _ => None
@@ -1497,20 +1823,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (if (let (x2, _) := idc_args in x2) <? 0
- then
- Some
- (UnderLet
- (#(Z_sub_get_borrow)%expr @ x @ x0 @
- (##(- (let (x2, _) := idc_args in x2))%Z)%expr)%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($vc)%expr,
- (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype ((ℤ -> ℤ) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ fv <- (if (let (x2, _) := idc_args in x2) <? 0
+ then
+ Some
+ (UnderLet
+ (#(Z_sub_get_borrow)%expr @ x @ x0 @
+ (##(- (let (x2, _) := idc_args in x2))%Z)%expr)%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($vc)%expr,
+ (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
| None => None
end
| _ => None
@@ -1524,11 +1855,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x2 <- (if (let (x2, _) := idc_args in x2) =? 0
- then Some (x1, (##0)%expr)%expr_pat
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype ((ℤ -> (projT1 args)) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x2 <- (if (let (x2, _) := idc_args in x2) =? 0
+ then Some (x1, (##0)%expr)%expr_pat
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -1542,11 +1878,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x2 <- (if (let (x2, _) := idc_args in x2) =? 0
- then Some (x0, (##0)%expr)%expr_pat
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype ((ℤ -> ℤ) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x2 <- (if (let (x2, _) := idc_args in x2) =? 0
+ then Some (x0, (##0)%expr)%expr_pat
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -1556,11 +1897,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((ℤ -> ℤ) -> ℤ)%ptype option (fun x2 : option => x2)
with
| Some (_, _, _) =>
- Some
- (UnderLet (#(Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($v)%expr, #(snd)%expr @ ($v)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq ((ℤ -> ℤ) -> ℤ)%ptype
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ then
+ Some
+ (UnderLet (#(Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
+ (fun v : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($v)%expr, #(snd)%expr @ ($v)%expr)%expr_pat))
+ else None
| None => None
end);;
None);;;
@@ -1576,11 +1922,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x2 <- (if (let (x2, _) := idc_args in x2) =? 0
- then Some (x0 + x1)%expr
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype (((projT1 args) -> ℤ) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x2 <- (if (let (x2, _) := idc_args in x2) =? 0
+ then Some (x0 + x1)%expr
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -1590,9 +1941,14 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((ℤ -> ℤ) -> ℤ)%ptype option (fun x2 : option => x2)
with
| Some (_, _, _) =>
- Some
- (UnderLet (#(Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var ℤ => Base ($v)%expr))
+ if
+ type.type_beq base.type base.type.type_beq ((ℤ -> ℤ) -> ℤ)%ptype
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ then
+ Some
+ (UnderLet (#(Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
+ (fun v : var ℤ => Base ($v)%expr))
+ else None
| None => None
end);;
None);;;
@@ -1612,15 +1968,54 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x4 : option => x4)
with
| Some (_, _, _, _) =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> (projT1 args0)) -> s) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ v <- type.try_make_transport_cps s ℤ;
+ fv <- (if (let (x4, _) := idc_args in x4) =? 0
+ then
+ Some
+ (UnderLet
+ (#(Z_sub_get_borrow)%expr @ x @ x2 @
+ v (Compile.reflect x3))%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($vc)%expr,
+ (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end);;
+ _ <- invert_bind_args idc0 Raw.ident.Z_opp;
+ args0 <- invert_bind_args idc Raw.ident.Literal;
+ match
+ pattern.type.unify_extracted_cps
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> (projT1 args0)) -> s) -> ℤ)%ptype option
+ (fun x4 : option => x4)
+ with
+ | Some (_, _, _, _) =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> (projT1 args0)) -> s) -> ℤ)%ptype
+ then
idc_args <- ident.unify pattern.ident.Literal
##(projT2 args0);
v <- type.try_make_transport_cps s ℤ;
- fv <- (if (let (x4, _) := idc_args in x4) =? 0
+ fv <- (if (let (x4, _) := idc_args in x4) <? 0
then
Some
(UnderLet
- (#(Z_sub_get_borrow)%expr @ x @ x2 @
- v (Compile.reflect x3))%expr_pat
+ (#(Z_sub_with_get_borrow)%expr @ x @
+ (##(- (let (x4, _) := idc_args in x4))%Z)%expr @
+ x2 @ v (Compile.reflect x3))%expr_pat
(fun vc : var (ℤ * ℤ)%etype =>
Base
(#(fst)%expr @ ($vc)%expr,
@@ -1628,34 +2023,7 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
else None);
Some (fv0 <-- fv;
Base fv0)%under_lets
- | None => None
- end);;
- _ <- invert_bind_args idc0 Raw.ident.Z_opp;
- args0 <- invert_bind_args idc Raw.ident.Literal;
- match
- pattern.type.unify_extracted_cps
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((ℤ -> (projT1 args0)) -> s) -> ℤ)%ptype option
- (fun x4 : option => x4)
- with
- | Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s ℤ;
- fv <- (if (let (x4, _) := idc_args in x4) <? 0
- then
- Some
- (UnderLet
- (#(Z_sub_with_get_borrow)%expr @ x @
- (##(- (let (x4, _) := idc_args in x4))%Z)%expr @
- x2 @ v (Compile.reflect x3))%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($vc)%expr,
- (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -1674,15 +2042,54 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x4 : option => x4)
with
| Some (_, _, _, _) =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> (projT1 args0)) -> ℤ) -> s)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ v <- type.try_make_transport_cps s ℤ;
+ fv <- (if (let (x4, _) := idc_args in x4) =? 0
+ then
+ Some
+ (UnderLet
+ (#(Z_sub_get_borrow)%expr @ x @ x1 @
+ v (Compile.reflect x3))%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($vc)%expr,
+ (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end);;
+ _ <- invert_bind_args idc0 Raw.ident.Z_opp;
+ args0 <- invert_bind_args idc Raw.ident.Literal;
+ match
+ pattern.type.unify_extracted_cps
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> (projT1 args0)) -> ℤ) -> s)%ptype option
+ (fun x4 : option => x4)
+ with
+ | Some (_, _, _, _) =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> (projT1 args0)) -> ℤ) -> s)%ptype
+ then
idc_args <- ident.unify pattern.ident.Literal
##(projT2 args0);
v <- type.try_make_transport_cps s ℤ;
- fv <- (if (let (x4, _) := idc_args in x4) =? 0
+ fv <- (if (let (x4, _) := idc_args in x4) <? 0
then
Some
(UnderLet
- (#(Z_sub_get_borrow)%expr @ x @ x1 @
- v (Compile.reflect x3))%expr_pat
+ (#(Z_sub_with_get_borrow)%expr @ x @
+ (##(- (let (x4, _) := idc_args in x4))%Z)%expr @
+ x1 @ v (Compile.reflect x3))%expr_pat
(fun vc : var (ℤ * ℤ)%etype =>
Base
(#(fst)%expr @ ($vc)%expr,
@@ -1690,34 +2097,7 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
else None);
Some (fv0 <-- fv;
Base fv0)%under_lets
- | None => None
- end);;
- _ <- invert_bind_args idc0 Raw.ident.Z_opp;
- args0 <- invert_bind_args idc Raw.ident.Literal;
- match
- pattern.type.unify_extracted_cps
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((ℤ -> (projT1 args0)) -> ℤ) -> s)%ptype option
- (fun x4 : option => x4)
- with
- | Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s ℤ;
- fv <- (if (let (x4, _) := idc_args in x4) <? 0
- then
- Some
- (UnderLet
- (#(Z_sub_with_get_borrow)%expr @ x @
- (##(- (let (x4, _) := idc_args in x4))%Z)%expr @
- x1 @ v (Compile.reflect x3))%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($vc)%expr,
- (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
+ else None
| None => None
end
| @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
@@ -1736,29 +2116,35 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (if
- ((let (x3, _) := idc_args0 in x3) <=? 0) &&
- ((let (x3, _) := idc_args in x3) <=? 0) &&
- ((let (x3, _) := idc_args0 in x3) +
- (let (x3, _) := idc_args in x3) <? 0)
- then
- Some
- (UnderLet
- (#(Z_sub_with_get_borrow)%expr @ x @
- (##(- (let (x3, _) := idc_args in x3))%Z)%expr @
- x2 @
- (##(- (let (x3, _) := idc_args0 in x3))%Z)%expr)%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($vc)%expr,
- (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> (projT1 args0)) -> (projT1 args)) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (if
+ ((let (x3, _) := idc_args0 in x3) <=? 0) &&
+ ((let (x3, _) := idc_args in x3) <=? 0) &&
+ ((let (x3, _) := idc_args0 in x3) +
+ (let (x3, _) := idc_args in x3) <? 0)
+ then
+ Some
+ (UnderLet
+ (#(Z_sub_with_get_borrow)%expr @ x @
+ (##(- (let (x3, _) := idc_args in x3))%Z)%expr @
+ x2 @
+ (##(- (let (x3, _) := idc_args0 in x3))%Z)%expr)%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($vc)%expr,
+ (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
| None => None
end
| _ => None
@@ -1774,29 +2160,35 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (if
- ((let (x3, _) := idc_args0 in x3) <=? 0) &&
- ((let (x3, _) := idc_args in x3) <=? 0) &&
- ((let (x3, _) := idc_args0 in x3) +
- (let (x3, _) := idc_args in x3) <? 0)
- then
- Some
- (UnderLet
- (#(Z_sub_with_get_borrow)%expr @ x @
- (##(- (let (x3, _) := idc_args in x3))%Z)%expr @
- x1 @
- (##(- (let (x3, _) := idc_args0 in x3))%Z)%expr)%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($vc)%expr,
- (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> (projT1 args0)) -> ℤ) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (if
+ ((let (x3, _) := idc_args0 in x3) <=? 0) &&
+ ((let (x3, _) := idc_args in x3) <=? 0) &&
+ ((let (x3, _) := idc_args0 in x3) +
+ (let (x3, _) := idc_args in x3) <? 0)
+ then
+ Some
+ (UnderLet
+ (#(Z_sub_with_get_borrow)%expr @ x @
+ (##(- (let (x3, _) := idc_args in x3))%Z)%expr @
+ x1 @
+ (##(- (let (x3, _) := idc_args0 in x3))%Z)%expr)%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($vc)%expr,
+ (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
| None => None
end
| _ => None
@@ -1815,18 +2207,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
ℤ)%ptype option (fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- _ <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x3 <- (if
- ((let (x3, _) := idc_args0 in x3) =? 0) &&
- ((let (x3, _) := idc_args1 in x3) =? 0)
- then Some (x2, (##0)%expr)%expr_pat
- else None);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args0) -> (projT1 args1)) ->
+ (projT1 args)) -> ℤ)%ptype
+ then
+ _ <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x3 <- (if
+ ((let (x3, _) := idc_args0 in x3) =? 0) &&
+ ((let (x3, _) := idc_args1 in x3) =? 0)
+ then Some (x2, (##0)%expr)%expr_pat
+ else None);
+ Some (Base x3)
+ else None
| None => None
end
| _ => None
@@ -1843,18 +2242,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(projT1 args))%ptype option (fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- _ <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x3 <- (if
- ((let (x3, _) := idc_args0 in x3) =? 0) &&
- ((let (x3, _) := idc_args1 in x3) =? 0)
- then Some (x1, (##0)%expr)%expr_pat
- else None);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args0) -> (projT1 args1)) -> ℤ) ->
+ (projT1 args))%ptype
+ then
+ _ <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x3 <- (if
+ ((let (x3, _) := idc_args0 in x3) =? 0) &&
+ ((let (x3, _) := idc_args1 in x3) =? 0)
+ then Some (x1, (##0)%expr)%expr_pat
+ else None);
+ Some (Base x3)
+ else None
| None => None
end
| _ => None
@@ -1868,19 +2274,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (if (let (x3, _) := idc_args in x3) =? 0
- then
- Some
- (UnderLet
- (#(Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($vc)%expr,
- #(snd)%expr @ ($vc)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> (projT1 args)) -> ℤ) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ fv <- (if (let (x3, _) := idc_args in x3) =? 0
+ then
+ Some
+ (UnderLet
+ (#(Z_add_get_carry)%expr @ x @ x1 @ x2)%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($vc)%expr,
+ #(snd)%expr @ ($vc)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x3 =>
@@ -1894,17 +2306,23 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, _, _) =>
- v <- type.try_make_transport_cps s ℤ;
- v0 <- type.try_make_transport_cps s0 ℤ;
- Some
- (UnderLet
- (#(Z_sub_with_get_borrow)%expr @ x @
- v (Compile.reflect x3) @ x2 @
- v0 (Compile.reflect x4))%expr_pat
- (fun v1 : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($v1)%expr,
- (- (#(snd)%expr @ $v1)%expr_pat)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> s) -> s0) -> ℤ)%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ v0 <- type.try_make_transport_cps s0 ℤ;
+ Some
+ (UnderLet
+ (#(Z_sub_with_get_borrow)%expr @ x @
+ v (Compile.reflect x3) @ x2 @
+ v0 (Compile.reflect x4))%expr_pat
+ (fun v1 : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($v1)%expr,
+ (- (#(snd)%expr @ $v1)%expr_pat)%expr)%expr_pat))
+ else None
| None => None
end
| @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
@@ -1923,17 +2341,23 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, _, _) =>
- v <- type.try_make_transport_cps s ℤ;
- v0 <- type.try_make_transport_cps s0 ℤ;
- Some
- (UnderLet
- (#(Z_sub_with_get_borrow)%expr @ x @
- v (Compile.reflect x3) @ x1 @
- v0 (Compile.reflect x4))%expr_pat
- (fun v1 : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($v1)%expr,
- (- (#(snd)%expr @ $v1)%expr_pat)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> s) -> ℤ) -> s0)%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ v0 <- type.try_make_transport_cps s0 ℤ;
+ Some
+ (UnderLet
+ (#(Z_sub_with_get_borrow)%expr @ x @
+ v (Compile.reflect x3) @ x1 @
+ v0 (Compile.reflect x4))%expr_pat
+ (fun v1 : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($v1)%expr,
+ (- (#(snd)%expr @ $v1)%expr_pat)%expr)%expr_pat))
+ else None
| None => None
end
| @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
@@ -1953,23 +2377,29 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x4 : option => x4)
with
| Some (_, _, _, _) =>
- v <- type.try_make_transport_cps s ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (if (let (x4, _) := idc_args in x4) <=? 0
- then
- Some
- (UnderLet
- (#(Z_sub_with_get_borrow)%expr @ x @
- v (Compile.reflect x3) @ x2 @
- (##(- (let (x4, _) := idc_args in x4))%Z)%expr)%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($vc)%expr,
- (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> s) -> (projT1 args)) -> ℤ)%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (if (let (x4, _) := idc_args in x4) <=? 0
+ then
+ Some
+ (UnderLet
+ (#(Z_sub_with_get_borrow)%expr @ x @
+ v (Compile.reflect x3) @ x2 @
+ (##(- (let (x4, _) := idc_args in x4))%Z)%expr)%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($vc)%expr,
+ (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
| None => None
end
| _ => None
@@ -1985,23 +2415,29 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x4 : option => x4)
with
| Some (_, _, _, _) =>
- v <- type.try_make_transport_cps s ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (if (let (x4, _) := idc_args in x4) <=? 0
- then
- Some
- (UnderLet
- (#(Z_sub_with_get_borrow)%expr @ x @
- v (Compile.reflect x3) @ x1 @
- (##(- (let (x4, _) := idc_args in x4))%Z)%expr)%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($vc)%expr,
- (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> s) -> ℤ) -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (if (let (x4, _) := idc_args in x4) <=? 0
+ then
+ Some
+ (UnderLet
+ (#(Z_sub_with_get_borrow)%expr @ x @
+ v (Compile.reflect x3) @ x1 @
+ (##(- (let (x4, _) := idc_args in x4))%Z)%expr)%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($vc)%expr,
+ (- (#(snd)%expr @ $vc)%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
| None => None
end
| _ => None
@@ -2024,25 +2460,31 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (if
- ((let (x3, _) := idc_args in x3) =? 0) &&
- ((let (x3, _) := idc_args0 in x3) =? 0)
- then
- Some
- (UnderLet
- (#(Z_add_with_get_carry)%expr @ x @ x0 @
- (##(let (x3, _) := idc_args in x3))%expr @
- (##(let (x3, _) := idc_args0 in x3))%expr)%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($vc)%expr, (##0)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((ℤ -> ℤ) -> (projT1 args0)) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (if
+ ((let (x3, _) := idc_args in x3) =? 0) &&
+ ((let (x3, _) := idc_args0 in x3) =? 0)
+ then
+ Some
+ (UnderLet
+ (#(Z_add_with_get_carry)%expr @ x @ x0 @
+ (##(let (x3, _) := idc_args in x3))%expr @
+ (##(let (x3, _) := idc_args0 in x3))%expr)%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($vc)%expr, (##0)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
| None => None
end
| _ => None
@@ -2054,12 +2496,17 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype option (fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- Some
- (UnderLet
- (#(Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($v)%expr, #(snd)%expr @ ($v)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ then
+ Some
+ (UnderLet
+ (#(Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
+ (fun v : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($v)%expr, #(snd)%expr @ ($v)%expr)%expr_pat))
+ else None
| None => None
end);;
None);;;
@@ -2071,11 +2518,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((ℤ -> ℤ) -> ℤ)%ptype option (fun x2 : option => x2)
with
| Some (_, _, _) =>
- Some
- (UnderLet (#(Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
- (fun v : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($v)%expr, #(snd)%expr @ ($v)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq ((ℤ -> ℤ) -> ℤ)%ptype
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ then
+ Some
+ (UnderLet (#(Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
+ (fun v : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($v)%expr, #(snd)%expr @ ($v)%expr)%expr_pat))
+ else None
| None => None
end;;;
Base (#(Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat)%option
@@ -2086,12 +2538,17 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype option (fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- Some
- (UnderLet
- (#(Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
- (fun v : var (ℤ * ℤ)%etype =>
- Base
- (#(fst)%expr @ ($v)%expr, #(snd)%expr @ ($v)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ then
+ Some
+ (UnderLet
+ (#(Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
+ (fun v : var (ℤ * ℤ)%etype =>
+ Base
+ (#(fst)%expr @ ($v)%expr, #(snd)%expr @ ($v)%expr)%expr_pat))
+ else None
| None => None
end;;;
Base (#(Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat)%option
@@ -2120,16 +2577,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
s0) -> s)%ptype option (fun x2 : option => x2)
with
| Some (_, (_, (_, _)), _, _) =>
- _ <- ident.unify pattern.ident.pair pair;
- v <- type.try_make_transport_cps s0 ℤ;
- v0 <- type.try_make_transport_cps s ℤ;
- Some
- (fv <-- do_again (ℤ * ℤ)
- (#(Z_cast (Datatypes.fst range))%expr @
- ($(v (Compile.reflect x1)))%expr,
- #(Z_cast (Datatypes.snd range))%expr @
- ($(v0 (Compile.reflect x0)))%expr)%expr_pat;
- Base fv)%under_lets
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ -> (ℤ * ℤ)%etype) -> ℤ) -> ℤ)%ptype
+ ((((let (x2, _) := args in x2) ->
+ (let (_, y) := args in y) ->
+ ((let (x2, _) := args in x2) * (let (_, y) := args in y))%etype) ->
+ s0) -> s)%ptype
+ then
+ _ <- ident.unify pattern.ident.pair pair;
+ v <- type.try_make_transport_cps s0 ℤ;
+ v0 <- type.try_make_transport_cps s ℤ;
+ Some
+ (fv <-- do_again (ℤ * ℤ)
+ (#(Z_cast (Datatypes.fst range))%expr @
+ ($(v (Compile.reflect x1)))%expr,
+ #(Z_cast (Datatypes.snd range))%expr @
+ ($(v0 (Compile.reflect x0)))%expr)%expr_pat;
+ Base fv)%under_lets
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App _
diff --git a/src/Experiments/NewPipeline/fancy_rewrite_head.out b/src/Experiments/NewPipeline/fancy_rewrite_head.out
index 235b8ad75..d6ae37896 100644
--- a/src/Experiments/NewPipeline/fancy_rewrite_head.out
+++ b/src/Experiments/NewPipeline/fancy_rewrite_head.out
@@ -132,31 +132,37 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x3 <- (if
- (let (x3, _) := idc_args0 in x3) =?
- 2
- ^ (2 *
- Z.log2_up (let (x3, _) := idc_args0 in x3) /
- 2) - 1
- then
- x3 <- invert_low
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args1) -> s0 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x3 <- (if
+ (let (x3, _) := idc_args0 in x3) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x3, _) := idc_args0 in x3) /
+ 2) - 1
+ then
+ x3 <- invert_low
+ (2 *
+ Z.log2_up
+ (let (x3, _) := idc_args0 in x3))
+ (let (x3, _) := idc_args in x3);
+ Some
+ (#(fancy_mulll
(2 *
Z.log2_up
- (let (x3, _) := idc_args0 in x3))
- (let (x3, _) := idc_args in x3);
- Some
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let (x4, _) := idc_args0 in x4)))%expr @
- ((##x3)%expr, v (Compile.reflect x2)))%expr_pat
- else None);
- Some (Base x3)
+ (let (x4, _) := idc_args0 in x4)))%expr @
+ ((##x3)%expr, v (Compile.reflect x2)))%expr_pat
+ else None);
+ Some (Base x3)
+ else None
| None => None
end
| _ => None
@@ -172,6 +178,53 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, (_, _)) =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args1) -> (projT1 args) -> s)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v <- type.try_make_transport_cps s ℤ;
+ x3 <- (if
+ (let (x3, _) := idc_args0 in x3) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x3, _) := idc_args0 in x3) /
+ 2) - 1
+ then
+ x3 <- invert_low
+ (2 *
+ Z.log2_up
+ (let (x3, _) := idc_args0 in x3))
+ (let (x3, _) := idc_args in x3);
+ Some
+ (#(fancy_mulll
+ (2 *
+ Z.log2_up
+ (let (x4, _) := idc_args0 in x4)))%expr @
+ ((##x3)%expr, v (Compile.reflect x1)))%expr_pat
+ else None);
+ Some (Base x3)
+ else None
+ | None => None
+ end);;
+ args <- invert_bind_args idc1 Raw.ident.Literal;
+ _ <- invert_bind_args idc0 Raw.ident.Z_land;
+ args1 <- invert_bind_args idc Raw.ident.Literal;
+ match
+ pattern.type.unify_extracted_cps (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args1) -> (projT1 args) -> s)%ptype option
+ (fun x3 : option => x3)
+ with
+ | Some (_, (_, _)) =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args1) -> (projT1 args) -> s)%ptype
+ then
idc_args <- ident.unify pattern.ident.Literal
##(projT2 args1);
idc_args0 <- ident.unify pattern.ident.Literal
@@ -184,55 +237,20 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
Z.log2_up (let (x3, _) := idc_args0 in x3) /
2) - 1
then
- x3 <- invert_low
+ x3 <- invert_high
(2 *
Z.log2_up
(let (x3, _) := idc_args0 in x3))
(let (x3, _) := idc_args in x3);
Some
- (#(fancy_mulll
+ (#(fancy_mulhl
(2 *
Z.log2_up
(let (x4, _) := idc_args0 in x4)))%expr @
((##x3)%expr, v (Compile.reflect x1)))%expr_pat
else None);
Some (Base x3)
- | None => None
- end);;
- args <- invert_bind_args idc1 Raw.ident.Literal;
- _ <- invert_bind_args idc0 Raw.ident.Z_land;
- args1 <- invert_bind_args idc Raw.ident.Literal;
- match
- pattern.type.unify_extracted_cps (ℤ -> ℤ -> ℤ)%ptype
- ((projT1 args1) -> (projT1 args) -> s)%ptype option
- (fun x3 : option => x3)
- with
- | Some (_, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v <- type.try_make_transport_cps s ℤ;
- x3 <- (if
- (let (x3, _) := idc_args0 in x3) =?
- 2
- ^ (2 *
- Z.log2_up (let (x3, _) := idc_args0 in x3) /
- 2) - 1
- then
- x3 <- invert_high
- (2 *
- Z.log2_up
- (let (x3, _) := idc_args0 in x3))
- (let (x3, _) := idc_args in x3);
- Some
- (#(fancy_mulhl
- (2 *
- Z.log2_up
- (let (x4, _) := idc_args0 in x4)))%expr @
- ((##x3)%expr, v (Compile.reflect x1)))%expr_pat
- else None);
- Some (Base x3)
+ else None
| None => None
end
| _ => None
@@ -248,31 +266,37 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x3 <- (if
- (let (x3, _) := idc_args0 in x3) =?
- 2
- ^ (2 *
- Z.log2_up (let (x3, _) := idc_args0 in x3) /
- 2) - 1
- then
- x3 <- invert_high
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args1) -> s0 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x3 <- (if
+ (let (x3, _) := idc_args0 in x3) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x3, _) := idc_args0 in x3) /
+ 2) - 1
+ then
+ x3 <- invert_high
+ (2 *
+ Z.log2_up
+ (let (x3, _) := idc_args0 in x3))
+ (let (x3, _) := idc_args in x3);
+ Some
+ (#(fancy_mulhl
(2 *
Z.log2_up
- (let (x3, _) := idc_args0 in x3))
- (let (x3, _) := idc_args in x3);
- Some
- (#(fancy_mulhl
- (2 *
- Z.log2_up
- (let (x4, _) := idc_args0 in x4)))%expr @
- ((##x3)%expr, v (Compile.reflect x2)))%expr_pat
- else None);
- Some (Base x3)
+ (let (x4, _) := idc_args0 in x4)))%expr @
+ ((##x3)%expr, v (Compile.reflect x2)))%expr_pat
+ else None);
+ Some (Base x3)
+ else None
| None => None
end
| _ => None
@@ -288,19 +312,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x3 <- (x3 <- invert_low
- (2 * (let (x3, _) := idc_args0 in x3))
- (let (x3, _) := idc_args in x3);
- Some
- (#(fancy_mullh
- (2 * (let (x4, _) := idc_args0 in x4)))%expr @
- ((##x3)%expr, v (Compile.reflect x2)))%expr_pat);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args1) -> s0 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x3 <- (x3 <- invert_low
+ (2 * (let (x3, _) := idc_args0 in x3))
+ (let (x3, _) := idc_args in x3);
+ Some
+ (#(fancy_mullh
+ (2 * (let (x4, _) := idc_args0 in x4)))%expr @
+ ((##x3)%expr, v (Compile.reflect x2)))%expr_pat);
+ Some (Base x3)
+ else None
| None => None
end);;
args <- invert_bind_args idc1 Raw.ident.Literal;
@@ -312,19 +342,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x3 <- (x3 <- invert_high
- (2 * (let (x3, _) := idc_args0 in x3))
- (let (x3, _) := idc_args in x3);
- Some
- (#(fancy_mulhh
- (2 * (let (x4, _) := idc_args0 in x4)))%expr @
- ((##x3)%expr, v (Compile.reflect x2)))%expr_pat);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args1) -> s0 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x3 <- (x3 <- invert_high
+ (2 * (let (x3, _) := idc_args0 in x3))
+ (let (x3, _) := idc_args in x3);
+ Some
+ (#(fancy_mulhh
+ (2 * (let (x4, _) := idc_args0 in x4)))%expr @
+ ((##x3)%expr, v (Compile.reflect x2)))%expr_pat);
+ Some (Base x3)
+ else None
| None => None
end
| _ => None
@@ -355,31 +391,37 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x3 <- (if
- (let (x3, _) := idc_args in x3) =?
- 2
- ^ (2 *
- Z.log2_up (let (x3, _) := idc_args in x3) /
- 2) - 1
- then
- y <- invert_low
- (2 *
- Z.log2_up
- (let (x3, _) := idc_args in x3))
- (let (x3, _) := idc_args0 in x3);
- Some
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let (x3, _) := idc_args in x3)))%expr @
- (v (Compile.reflect x1), (##y)%expr))%expr_pat
- else None);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args0) -> s) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ v <- type.try_make_transport_cps s ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x3 <- (if
+ (let (x3, _) := idc_args in x3) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x3, _) := idc_args in x3) /
+ 2) - 1
+ then
+ y <- invert_low
+ (2 *
+ Z.log2_up
+ (let (x3, _) := idc_args in x3))
+ (let (x3, _) := idc_args0 in x3);
+ Some
+ (#(fancy_mulll
+ (2 *
+ Z.log2_up
+ (let (x3, _) := idc_args in x3)))%expr @
+ (v (Compile.reflect x1), (##y)%expr))%expr_pat
+ else None);
+ Some (Base x3)
+ else None
| None => None
end
| _ => None
@@ -399,31 +441,37 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, _, _) =>
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x3 <- (if
- (let (x3, _) := idc_args in x3) =?
- 2
- ^ (2 *
- Z.log2_up (let (x3, _) := idc_args in x3) /
- 2) - 1
- then
- y <- invert_low
- (2 *
- Z.log2_up
- (let (x3, _) := idc_args in x3))
- (let (x3, _) := idc_args0 in x3);
- Some
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let (x3, _) := idc_args in x3)))%expr @
- (v (Compile.reflect x2), (##y)%expr))%expr_pat
- else None);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ ((s0 -> (projT1 args0)) -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x3 <- (if
+ (let (x3, _) := idc_args in x3) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x3, _) := idc_args in x3) /
+ 2) - 1
+ then
+ y <- invert_low
+ (2 *
+ Z.log2_up
+ (let (x3, _) := idc_args in x3))
+ (let (x3, _) := idc_args0 in x3);
+ Some
+ (#(fancy_mulll
+ (2 *
+ Z.log2_up
+ (let (x3, _) := idc_args in x3)))%expr @
+ (v (Compile.reflect x2), (##y)%expr))%expr_pat
+ else None);
+ Some (Base x3)
+ else None
| None => None
end
| _ => None
@@ -443,31 +491,37 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x3 <- (if
- (let (x3, _) := idc_args in x3) =?
- 2
- ^ (2 *
- Z.log2_up (let (x3, _) := idc_args in x3) /
- 2) - 1
- then
- y <- invert_high
- (2 *
- Z.log2_up
- (let (x3, _) := idc_args in x3))
- (let (x3, _) := idc_args0 in x3);
- Some
- (#(fancy_mullh
- (2 *
- Z.log2_up
- (let (x3, _) := idc_args in x3)))%expr @
- (v (Compile.reflect x1), (##y)%expr))%expr_pat
- else None);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args0) -> s) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ v <- type.try_make_transport_cps s ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x3 <- (if
+ (let (x3, _) := idc_args in x3) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x3, _) := idc_args in x3) /
+ 2) - 1
+ then
+ y <- invert_high
+ (2 *
+ Z.log2_up
+ (let (x3, _) := idc_args in x3))
+ (let (x3, _) := idc_args0 in x3);
+ Some
+ (#(fancy_mullh
+ (2 *
+ Z.log2_up
+ (let (x3, _) := idc_args in x3)))%expr @
+ (v (Compile.reflect x1), (##y)%expr))%expr_pat
+ else None);
+ Some (Base x3)
+ else None
| None => None
end
| _ => None
@@ -487,31 +541,37 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, _, _) =>
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x3 <- (if
- (let (x3, _) := idc_args in x3) =?
- 2
- ^ (2 *
- Z.log2_up (let (x3, _) := idc_args in x3) /
- 2) - 1
- then
- y <- invert_high
- (2 *
- Z.log2_up
- (let (x3, _) := idc_args in x3))
- (let (x3, _) := idc_args0 in x3);
- Some
- (#(fancy_mullh
- (2 *
- Z.log2_up
- (let (x3, _) := idc_args in x3)))%expr @
- (v (Compile.reflect x2), (##y)%expr))%expr_pat
- else None);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ ((s0 -> (projT1 args0)) -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x3 <- (if
+ (let (x3, _) := idc_args in x3) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x3, _) := idc_args in x3) /
+ 2) - 1
+ then
+ y <- invert_high
+ (2 *
+ Z.log2_up
+ (let (x3, _) := idc_args in x3))
+ (let (x3, _) := idc_args0 in x3);
+ Some
+ (#(fancy_mullh
+ (2 *
+ Z.log2_up
+ (let (x3, _) := idc_args in x3)))%expr @
+ (v (Compile.reflect x2), (##y)%expr))%expr_pat
+ else None);
+ Some (Base x3)
+ else None
| None => None
end
| _ => None
@@ -535,33 +595,39 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v0 <- type.try_make_transport_cps s1 ℤ;
- x5 <- (if
- ((let (x5, _) := idc_args in x5) =?
- 2
- ^ (2 *
- Z.log2_up (let (x5, _) := idc_args in x5) /
- 2) - 1) &&
- ((let (x5, _) := idc_args0 in x5) =?
- 2
- ^ (2 *
- Z.log2_up (let (x5, _) := idc_args in x5) /
- 2) - 1)
- then
- Some
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let (x5, _) := idc_args in x5)))%expr @
- (v (Compile.reflect x1),
- v0 (Compile.reflect x3)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ (((projT1 args1) -> s) -> (projT1 args) -> s1)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v0 <- type.try_make_transport_cps s1 ℤ;
+ x5 <- (if
+ ((let (x5, _) := idc_args in x5) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x5, _) := idc_args in x5) /
+ 2) - 1) &&
+ ((let (x5, _) := idc_args0 in x5) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x5, _) := idc_args in x5) /
+ 2) - 1)
+ then
+ Some
+ (#(fancy_mulll
+ (2 *
+ Z.log2_up
+ (let (x5, _) := idc_args in x5)))%expr @
+ (v (Compile.reflect x1),
+ v0 (Compile.reflect x3)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| @expr.App _ _ _ s1 _
@@ -601,33 +667,39 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, (_, _)) =>
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v0 <- type.try_make_transport_cps s1 ℤ;
- x5 <- (if
- ((let (x5, _) := idc_args in x5) =?
- 2
- ^ (2 *
- Z.log2_up (let (x5, _) := idc_args in x5) /
- 2) - 1) &&
- ((let (x5, _) := idc_args0 in x5) =?
- 2
- ^ (2 *
- Z.log2_up (let (x5, _) := idc_args in x5) /
- 2) - 1)
- then
- Some
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let (x5, _) := idc_args in x5)))%expr @
- (v (Compile.reflect x2),
- v0 (Compile.reflect x3)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s0 -> (projT1 args1)) -> (projT1 args) -> s1)%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v0 <- type.try_make_transport_cps s1 ℤ;
+ x5 <- (if
+ ((let (x5, _) := idc_args in x5) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x5, _) := idc_args in x5) /
+ 2) - 1) &&
+ ((let (x5, _) := idc_args0 in x5) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x5, _) := idc_args in x5) /
+ 2) - 1)
+ then
+ Some
+ (#(fancy_mulll
+ (2 *
+ Z.log2_up
+ (let (x5, _) := idc_args in x5)))%expr @
+ (v (Compile.reflect x2),
+ v0 (Compile.reflect x3)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| @expr.App _ _ _ s1 _
@@ -666,33 +738,39 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s ℤ;
- v0 <- type.try_make_transport_cps s2 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x5 <- (if
- ((let (x5, _) := idc_args in x5) =?
- 2
- ^ (2 *
- Z.log2_up (let (x5, _) := idc_args in x5) /
- 2) - 1) &&
- ((let (x5, _) := idc_args0 in x5) =?
- 2
- ^ (2 *
- Z.log2_up (let (x5, _) := idc_args in x5) /
- 2) - 1)
- then
- Some
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let (x5, _) := idc_args in x5)))%expr @
- (v (Compile.reflect x1),
- v0 (Compile.reflect x4)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ (((projT1 args1) -> s) -> s2 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s ℤ;
+ v0 <- type.try_make_transport_cps s2 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x5 <- (if
+ ((let (x5, _) := idc_args in x5) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x5, _) := idc_args in x5) /
+ 2) - 1) &&
+ ((let (x5, _) := idc_args0 in x5) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x5, _) := idc_args in x5) /
+ 2) - 1)
+ then
+ Some
+ (#(fancy_mulll
+ (2 *
+ Z.log2_up
+ (let (x5, _) := idc_args in x5)))%expr @
+ (v (Compile.reflect x1),
+ v0 (Compile.reflect x4)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| (@expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t1 idc1) x4 @
@@ -728,33 +806,39 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, (_, _)) =>
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v0 <- type.try_make_transport_cps s2 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x5 <- (if
- ((let (x5, _) := idc_args in x5) =?
- 2
- ^ (2 *
- Z.log2_up (let (x5, _) := idc_args in x5) /
- 2) - 1) &&
- ((let (x5, _) := idc_args0 in x5) =?
- 2
- ^ (2 *
- Z.log2_up (let (x5, _) := idc_args in x5) /
- 2) - 1)
- then
- Some
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let (x5, _) := idc_args in x5)))%expr @
- (v (Compile.reflect x2),
- v0 (Compile.reflect x4)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s0 -> (projT1 args1)) -> s2 -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v0 <- type.try_make_transport_cps s2 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x5 <- (if
+ ((let (x5, _) := idc_args in x5) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x5, _) := idc_args in x5) /
+ 2) - 1) &&
+ ((let (x5, _) := idc_args0 in x5) =?
+ 2
+ ^ (2 *
+ Z.log2_up (let (x5, _) := idc_args in x5) /
+ 2) - 1)
+ then
+ Some
+ (#(fancy_mulll
+ (2 *
+ Z.log2_up
+ (let (x5, _) := idc_args in x5)))%expr @
+ (v (Compile.reflect x2),
+ v0 (Compile.reflect x4)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| (@expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t1 idc1) x4 @
@@ -790,24 +874,30 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s ℤ;
- v0 <- type.try_make_transport_cps s2 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x5 <- (if
- (let (x5, _) := idc_args in x5) =?
- 2 ^ (2 * (let (x5, _) := idc_args0 in x5) / 2) -
- 1
- then
- Some
- (#(fancy_mullh
- (2 * (let (x5, _) := idc_args0 in x5)))%expr @
- (v (Compile.reflect x1),
- v0 (Compile.reflect x4)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ (((projT1 args1) -> s) -> s2 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s ℤ;
+ v0 <- type.try_make_transport_cps s2 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x5 <- (if
+ (let (x5, _) := idc_args in x5) =?
+ 2 ^ (2 * (let (x5, _) := idc_args0 in x5) / 2) -
+ 1
+ then
+ Some
+ (#(fancy_mullh
+ (2 * (let (x5, _) := idc_args0 in x5)))%expr @
+ (v (Compile.reflect x1),
+ v0 (Compile.reflect x4)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| (@expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t1 idc1) x4 @
@@ -843,24 +933,30 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, (_, _)) =>
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v0 <- type.try_make_transport_cps s2 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x5 <- (if
- (let (x5, _) := idc_args in x5) =?
- 2 ^ (2 * (let (x5, _) := idc_args0 in x5) / 2) -
- 1
- then
- Some
- (#(fancy_mullh
- (2 * (let (x5, _) := idc_args0 in x5)))%expr @
- (v (Compile.reflect x2),
- v0 (Compile.reflect x4)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s0 -> (projT1 args1)) -> s2 -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v0 <- type.try_make_transport_cps s2 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x5 <- (if
+ (let (x5, _) := idc_args in x5) =?
+ 2 ^ (2 * (let (x5, _) := idc_args0 in x5) / 2) -
+ 1
+ then
+ Some
+ (#(fancy_mullh
+ (2 * (let (x5, _) := idc_args0 in x5)))%expr @
+ (v (Compile.reflect x2),
+ v0 (Compile.reflect x4)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| (@expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t1 idc1) x4 @
@@ -893,19 +989,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, _, _) =>
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x3 <- (y <- invert_low
- (2 * (let (x3, _) := idc_args in x3))
- (let (x3, _) := idc_args0 in x3);
- Some
- (#(fancy_mulhl
- (2 * (let (x3, _) := idc_args in x3)))%expr @
- (v (Compile.reflect x2), (##y)%expr))%expr_pat);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ ((s0 -> (projT1 args0)) -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x3 <- (y <- invert_low
+ (2 * (let (x3, _) := idc_args in x3))
+ (let (x3, _) := idc_args0 in x3);
+ Some
+ (#(fancy_mulhl
+ (2 * (let (x3, _) := idc_args in x3)))%expr @
+ (v (Compile.reflect x2), (##y)%expr))%expr_pat);
+ Some (Base x3)
+ else None
| None => None
end);;
args <- invert_bind_args idc1 Raw.ident.Literal;
@@ -917,19 +1019,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, _, _) =>
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x3 <- (y <- invert_high
- (2 * (let (x3, _) := idc_args in x3))
- (let (x3, _) := idc_args0 in x3);
- Some
- (#(fancy_mulhh
- (2 * (let (x3, _) := idc_args in x3)))%expr @
- (v (Compile.reflect x2), (##y)%expr))%expr_pat);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ ((s0 -> (projT1 args0)) -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x3 <- (y <- invert_high
+ (2 * (let (x3, _) := idc_args in x3))
+ (let (x3, _) := idc_args0 in x3);
+ Some
+ (#(fancy_mulhh
+ (2 * (let (x3, _) := idc_args in x3)))%expr @
+ (v (Compile.reflect x2), (##y)%expr))%expr_pat);
+ Some (Base x3)
+ else None
| None => None
end
| @expr.App _ _ _ s1 _
@@ -947,25 +1055,32 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, (_, _)) =>
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v0 <- type.try_make_transport_cps s1 ℤ;
- x5 <- (if
- (let (x5, _) := idc_args0 in x5) =?
- 2
- ^ (2 * (let (x5, _) := idc_args in x5) / 2) -
- 1
- then
- Some
- (#(fancy_mulhl
- (2 * (let (x5, _) := idc_args in x5)))%expr @
- (v (Compile.reflect x2),
- v0 (Compile.reflect x3)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s0 -> (projT1 args1)) -> (projT1 args) -> s1)%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v0 <- type.try_make_transport_cps s1 ℤ;
+ x5 <- (if
+ (let (x5, _) := idc_args0 in x5) =?
+ 2
+ ^ (2 * (let (x5, _) := idc_args in x5) / 2) -
+ 1
+ then
+ Some
+ (#(fancy_mulhl
+ (2 *
+ (let (x5, _) := idc_args in x5)))%expr @
+ (v (Compile.reflect x2),
+ v0 (Compile.reflect x3)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| _ => None
@@ -983,25 +1098,32 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, (_, _)) =>
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v0 <- type.try_make_transport_cps s2 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x5 <- (if
- (let (x5, _) := idc_args0 in x5) =?
- 2
- ^ (2 * (let (x5, _) := idc_args in x5) / 2) -
- 1
- then
- Some
- (#(fancy_mulhl
- (2 * (let (x5, _) := idc_args in x5)))%expr @
- (v (Compile.reflect x2),
- v0 (Compile.reflect x4)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s0 -> (projT1 args1)) -> s2 -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v0 <- type.try_make_transport_cps s2 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x5 <- (if
+ (let (x5, _) := idc_args0 in x5) =?
+ 2
+ ^ (2 * (let (x5, _) := idc_args in x5) / 2) -
+ 1
+ then
+ Some
+ (#(fancy_mulhl
+ (2 *
+ (let (x5, _) := idc_args in x5)))%expr @
+ (v (Compile.reflect x2),
+ v0 (Compile.reflect x4)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| _ => None
@@ -1019,23 +1141,29 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, (_, _)) =>
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v0 <- type.try_make_transport_cps s2 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x5 <- (if
- (let (x5, _) := idc_args in x5) =?
- (let (x5, _) := idc_args0 in x5)
- then
- Some
- (#(fancy_mulhh
- (2 * (let (x5, _) := idc_args in x5)))%expr @
- (v (Compile.reflect x2),
- v0 (Compile.reflect x4)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s0 -> (projT1 args1)) -> s2 -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v0 <- type.try_make_transport_cps s2 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x5 <- (if
+ (let (x5, _) := idc_args in x5) =?
+ (let (x5, _) := idc_args0 in x5)
+ then
+ Some
+ (#(fancy_mulhh
+ (2 * (let (x5, _) := idc_args in x5)))%expr @
+ (v (Compile.reflect x2),
+ v0 (Compile.reflect x4)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| _ => None
@@ -1100,22 +1228,28 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x4 : option => x4)
with
| Some (_, _, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x4 <- (if
- (let (x4, _) := idc_args in x4) =?
- 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
- then
- Some
- (#(fancy_add
- (Z.log2 (let (x4, _) := idc_args in x4))
- (let (x4, _) := idc_args0 in x4))%expr @
- (x0, v (Compile.reflect x3)))%expr_pat
- else None);
- Some (Base x4)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ (((projT1 args1) -> ℤ) -> s0 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x4 <- (if
+ (let (x4, _) := idc_args in x4) =?
+ 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
+ then
+ Some
+ (#(fancy_add
+ (Z.log2 (let (x4, _) := idc_args in x4))
+ (let (x4, _) := idc_args0 in x4))%expr @
+ (x0, v (Compile.reflect x3)))%expr_pat
+ else None);
+ Some (Base x4)
+ else None
| None => None
end
| (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x3 @ ($_)%expr)%expr_pat |
@@ -1143,22 +1277,28 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x4 : option => x4)
with
| Some (_, (_, _), _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x4 <- (if
- (let (x4, _) := idc_args in x4) =?
- 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
- then
- Some
- (#(fancy_add
- (Z.log2 (let (x4, _) := idc_args in x4))
- (let (x4, _) := idc_args0 in x4))%expr @
- (x1, v (Compile.reflect x3)))%expr_pat
- else None);
- Some (Base x4)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> s0 -> (projT1 args)) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x4 <- (if
+ (let (x4, _) := idc_args in x4) =?
+ 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
+ then
+ Some
+ (#(fancy_add
+ (Z.log2 (let (x4, _) := idc_args in x4))
+ (let (x4, _) := idc_args0 in x4))%expr @
+ (x1, v (Compile.reflect x3)))%expr_pat
+ else None);
+ Some (Base x4)
+ else None
| None => None
end
| (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x3 @ ($_)%expr)%expr_pat |
@@ -1186,22 +1326,28 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x4 : option => x4)
with
| Some (_, _, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x4 <- (if
- (let (x4, _) := idc_args in x4) =?
- 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
- then
- Some
- (#(fancy_add
- (Z.log2 (let (x4, _) := idc_args in x4))
- (- (let (x4, _) := idc_args0 in x4)))%expr @
- (x0, v (Compile.reflect x3)))%expr_pat
- else None);
- Some (Base x4)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ (((projT1 args1) -> ℤ) -> s0 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x4 <- (if
+ (let (x4, _) := idc_args in x4) =?
+ 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
+ then
+ Some
+ (#(fancy_add
+ (Z.log2 (let (x4, _) := idc_args in x4))
+ (- (let (x4, _) := idc_args0 in x4)))%expr @
+ (x0, v (Compile.reflect x3)))%expr_pat
+ else None);
+ Some (Base x4)
+ else None
| None => None
end
| (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x3 @ ($_)%expr)%expr_pat |
@@ -1229,22 +1375,28 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x4 : option => x4)
with
| Some (_, (_, _), _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x4 <- (if
- (let (x4, _) := idc_args in x4) =?
- 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
- then
- Some
- (#(fancy_add
- (Z.log2 (let (x4, _) := idc_args in x4))
- (- (let (x4, _) := idc_args0 in x4)))%expr @
- (x1, v (Compile.reflect x3)))%expr_pat
- else None);
- Some (Base x4)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> s0 -> (projT1 args)) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x4 <- (if
+ (let (x4, _) := idc_args in x4) =?
+ 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
+ then
+ Some
+ (#(fancy_add
+ (Z.log2 (let (x4, _) := idc_args in x4))
+ (- (let (x4, _) := idc_args0 in x4)))%expr @
+ (x1, v (Compile.reflect x3)))%expr_pat
+ else None);
+ Some (Base x4)
+ else None
| None => None
end
| (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x3 @ ($_)%expr)%expr_pat |
@@ -1267,16 +1419,22 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x2 <- (if
- (let (x2, _) := idc_args in x2) =?
- 2 ^ Z.log2 (let (x2, _) := idc_args in x2)
- then
- Some
- (#(fancy_add (Z.log2 (let (x2, _) := idc_args in x2))
- 0)%expr @ (x0, x1))%expr_pat
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype (((projT1 args) -> ℤ) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x2 <- (if
+ (let (x2, _) := idc_args in x2) =?
+ 2 ^ Z.log2 (let (x2, _) := idc_args in x2)
+ then
+ Some
+ (#(fancy_add
+ (Z.log2 (let (x2, _) := idc_args in x2)) 0)%expr @
+ (x0, x1))%expr_pat
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -1303,22 +1461,28 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, _, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x5 <- (if
- (let (x5, _) := idc_args in x5) =?
- 2 ^ Z.log2 (let (x5, _) := idc_args in x5)
- then
- Some
- (#(fancy_addc
- (Z.log2 (let (x5, _) := idc_args in x5))
- (let (x5, _) := idc_args0 in x5))%expr @
- (x0, x1, v (Compile.reflect x4)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
+ ((((projT1 args1) -> ℤ) -> ℤ) -> s0 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x5 <- (if
+ (let (x5, _) := idc_args in x5) =?
+ 2 ^ Z.log2 (let (x5, _) := idc_args in x5)
+ then
+ Some
+ (#(fancy_addc
+ (Z.log2 (let (x5, _) := idc_args in x5))
+ (let (x5, _) := idc_args0 in x5))%expr @
+ (x0, x1, v (Compile.reflect x4)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x4 @ ($_)%expr)%expr_pat |
@@ -1347,22 +1511,28 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, (_, _), _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x5 <- (if
- (let (x5, _) := idc_args in x5) =?
- 2 ^ Z.log2 (let (x5, _) := idc_args in x5)
- then
- Some
- (#(fancy_addc
- (Z.log2 (let (x5, _) := idc_args in x5))
- (let (x5, _) := idc_args0 in x5))%expr @
- (x0, x2, v (Compile.reflect x4)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ -> ℤ) -> ℤ)%ptype
+ ((((projT1 args1) -> ℤ) -> s0 -> (projT1 args)) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x5 <- (if
+ (let (x5, _) := idc_args in x5) =?
+ 2 ^ Z.log2 (let (x5, _) := idc_args in x5)
+ then
+ Some
+ (#(fancy_addc
+ (Z.log2 (let (x5, _) := idc_args in x5))
+ (let (x5, _) := idc_args0 in x5))%expr @
+ (x0, x2, v (Compile.reflect x4)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x4 @ ($_)%expr)%expr_pat |
@@ -1391,22 +1561,28 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, _, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x5 <- (if
- (let (x5, _) := idc_args in x5) =?
- 2 ^ Z.log2 (let (x5, _) := idc_args in x5)
- then
- Some
- (#(fancy_addc
- (Z.log2 (let (x5, _) := idc_args in x5))
- (- (let (x5, _) := idc_args0 in x5)))%expr @
- (x0, x1, v (Compile.reflect x4)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
+ ((((projT1 args1) -> ℤ) -> ℤ) -> s0 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x5 <- (if
+ (let (x5, _) := idc_args in x5) =?
+ 2 ^ Z.log2 (let (x5, _) := idc_args in x5)
+ then
+ Some
+ (#(fancy_addc
+ (Z.log2 (let (x5, _) := idc_args in x5))
+ (- (let (x5, _) := idc_args0 in x5)))%expr @
+ (x0, x1, v (Compile.reflect x4)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x4 @ ($_)%expr)%expr_pat |
@@ -1435,22 +1611,28 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, (_, _), _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x5 <- (if
- (let (x5, _) := idc_args in x5) =?
- 2 ^ Z.log2 (let (x5, _) := idc_args in x5)
- then
- Some
- (#(fancy_addc
- (Z.log2 (let (x5, _) := idc_args in x5))
- (- (let (x5, _) := idc_args0 in x5)))%expr @
- (x0, x2, v (Compile.reflect x4)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ -> ℤ) -> ℤ)%ptype
+ ((((projT1 args1) -> ℤ) -> s0 -> (projT1 args)) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x5 <- (if
+ (let (x5, _) := idc_args in x5) =?
+ 2 ^ Z.log2 (let (x5, _) := idc_args in x5)
+ then
+ Some
+ (#(fancy_addc
+ (Z.log2 (let (x5, _) := idc_args in x5))
+ (- (let (x5, _) := idc_args0 in x5)))%expr @
+ (x0, x2, v (Compile.reflect x4)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t0 idc0) x4 @ ($_)%expr)%expr_pat |
@@ -1473,17 +1655,23 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x3 <- (if
- (let (x3, _) := idc_args in x3) =?
- 2 ^ Z.log2 (let (x3, _) := idc_args in x3)
- then
- Some
- (#(fancy_addc
- (Z.log2 (let (x3, _) := idc_args in x3)) 0)%expr @
- (x0, x1, x2))%expr_pat
- else None);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args) -> ℤ) -> ℤ) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x3 <- (if
+ (let (x3, _) := idc_args in x3) =?
+ 2 ^ Z.log2 (let (x3, _) := idc_args in x3)
+ then
+ Some
+ (#(fancy_addc
+ (Z.log2 (let (x3, _) := idc_args in x3)) 0)%expr @
+ (x0, x1, x2))%expr_pat
+ else None);
+ Some (Base x3)
+ else None
| None => None
end
| _ => None
@@ -1508,22 +1696,28 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x4 : option => x4)
with
| Some (_, _, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x4 <- (if
- (let (x4, _) := idc_args in x4) =?
- 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
- then
- Some
- (#(fancy_sub
- (Z.log2 (let (x4, _) := idc_args in x4))
- (let (x4, _) := idc_args0 in x4))%expr @
- (x0, v (Compile.reflect x3)))%expr_pat
- else None);
- Some (Base x4)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ (((projT1 args1) -> ℤ) -> s0 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x4 <- (if
+ (let (x4, _) := idc_args in x4) =?
+ 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
+ then
+ Some
+ (#(fancy_sub
+ (Z.log2 (let (x4, _) := idc_args in x4))
+ (let (x4, _) := idc_args0 in x4))%expr @
+ (x0, v (Compile.reflect x3)))%expr_pat
+ else None);
+ Some (Base x4)
+ else None
| None => None
end
| _ => None
@@ -1540,22 +1734,28 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x4 : option => x4)
with
| Some (_, _, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x4 <- (if
- (let (x4, _) := idc_args in x4) =?
- 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
- then
- Some
- (#(fancy_sub
- (Z.log2 (let (x4, _) := idc_args in x4))
- (- (let (x4, _) := idc_args0 in x4)))%expr @
- (x0, v (Compile.reflect x3)))%expr_pat
- else None);
- Some (Base x4)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ (((projT1 args1) -> ℤ) -> s0 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x4 <- (if
+ (let (x4, _) := idc_args in x4) =?
+ 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
+ then
+ Some
+ (#(fancy_sub
+ (Z.log2 (let (x4, _) := idc_args in x4))
+ (- (let (x4, _) := idc_args0 in x4)))%expr @
+ (x0, v (Compile.reflect x3)))%expr_pat
+ else None);
+ Some (Base x4)
+ else None
| None => None
end
| _ => None
@@ -1574,16 +1774,22 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x2 <- (if
- (let (x2, _) := idc_args in x2) =?
- 2 ^ Z.log2 (let (x2, _) := idc_args in x2)
- then
- Some
- (#(fancy_sub (Z.log2 (let (x2, _) := idc_args in x2))
- 0)%expr @ (x0, x1))%expr_pat
- else None);
- Some (Base x2)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype (((projT1 args) -> ℤ) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x2 <- (if
+ (let (x2, _) := idc_args in x2) =?
+ 2 ^ Z.log2 (let (x2, _) := idc_args in x2)
+ then
+ Some
+ (#(fancy_sub
+ (Z.log2 (let (x2, _) := idc_args in x2)) 0)%expr @
+ (x0, x1))%expr_pat
+ else None);
+ Some (Base x2)
+ else None
| None => None
end
| _ => None
@@ -1608,22 +1814,28 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, _, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x5 <- (if
- (let (x5, _) := idc_args in x5) =?
- 2 ^ Z.log2 (let (x5, _) := idc_args in x5)
- then
- Some
- (#(fancy_subb
- (Z.log2 (let (x5, _) := idc_args in x5))
- (let (x5, _) := idc_args0 in x5))%expr @
- (x0, x1, v (Compile.reflect x4)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
+ ((((projT1 args1) -> ℤ) -> ℤ) -> s0 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x5 <- (if
+ (let (x5, _) := idc_args in x5) =?
+ 2 ^ Z.log2 (let (x5, _) := idc_args in x5)
+ then
+ Some
+ (#(fancy_subb
+ (Z.log2 (let (x5, _) := idc_args in x5))
+ (let (x5, _) := idc_args0 in x5))%expr @
+ (x0, x1, v (Compile.reflect x4)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| _ => None
@@ -1640,22 +1852,28 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x5 : option => x5)
with
| Some (_, _, _, (_, _)) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x5 <- (if
- (let (x5, _) := idc_args in x5) =?
- 2 ^ Z.log2 (let (x5, _) := idc_args in x5)
- then
- Some
- (#(fancy_subb
- (Z.log2 (let (x5, _) := idc_args in x5))
- (- (let (x5, _) := idc_args0 in x5)))%expr @
- (x0, x1, v (Compile.reflect x4)))%expr_pat
- else None);
- Some (Base x5)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
+ ((((projT1 args1) -> ℤ) -> ℤ) -> s0 -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x5 <- (if
+ (let (x5, _) := idc_args in x5) =?
+ 2 ^ Z.log2 (let (x5, _) := idc_args in x5)
+ then
+ Some
+ (#(fancy_subb
+ (Z.log2 (let (x5, _) := idc_args in x5))
+ (- (let (x5, _) := idc_args0 in x5)))%expr @
+ (x0, x1, v (Compile.reflect x4)))%expr_pat
+ else None);
+ Some (Base x5)
+ else None
| None => None
end
| _ => None
@@ -1674,17 +1892,23 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- x3 <- (if
- (let (x3, _) := idc_args in x3) =?
- 2 ^ Z.log2 (let (x3, _) := idc_args in x3)
- then
- Some
- (#(fancy_subb
- (Z.log2 (let (x3, _) := idc_args in x3)) 0)%expr @
- (x0, x1, x2))%expr_pat
- else None);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args) -> ℤ) -> ℤ) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ x3 <- (if
+ (let (x3, _) := idc_args in x3) =?
+ 2 ^ Z.log2 (let (x3, _) := idc_args in x3)
+ then
+ Some
+ (#(fancy_subb
+ (Z.log2 (let (x3, _) := idc_args in x3)) 0)%expr @
+ (x0, x1, x2))%expr_pat
+ else None);
+ Some (Base x3)
+ else None
| None => None
end
| _ => None
@@ -1707,19 +1931,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x4 : option => x4)
with
| Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v <- type.try_make_transport_cps s ℤ;
- x4 <- (if
- (let (x4, _) := idc_args in x4) =?
- 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
- then
- Some
- (#(fancy_selm
- (Z.log2 (let (x4, _) := idc_args in x4)))%expr @
- (v (Compile.reflect x2), x0, x1))%expr_pat
- else None);
- Some (Base x4)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args) -> s) -> ℤ) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v <- type.try_make_transport_cps s ℤ;
+ x4 <- (if
+ (let (x4, _) := idc_args in x4) =?
+ 2 ^ Z.log2 (let (x4, _) := idc_args in x4)
+ then
+ Some
+ (#(fancy_selm
+ (Z.log2 (let (x4, _) := idc_args in x4)))%expr @
+ (v (Compile.reflect x2), x0, x1))%expr_pat
+ else None);
+ Some (Base x4)
+ else None
| None => None
end
| _ => None
@@ -1735,16 +1965,22 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x4 : option => x4)
with
| Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v <- type.try_make_transport_cps s ℤ;
- x4 <- (if (let (x4, _) := idc_args in x4) =? 1
- then
- Some
- (#(fancy_sell)%expr @
- (v (Compile.reflect x2), x0, x1))%expr_pat
- else None);
- Some (Base x4)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args) -> s) -> ℤ) -> ℤ)%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v <- type.try_make_transport_cps s ℤ;
+ x4 <- (if (let (x4, _) := idc_args in x4) =? 1
+ then
+ Some
+ (#(fancy_sell)%expr @
+ (v (Compile.reflect x2), x0, x1))%expr_pat
+ else None);
+ Some (Base x4)
+ else None
| None => None
end
| _ => None
@@ -1760,16 +1996,22 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x4 : option => x4)
with
| Some (_, _, _, _) =>
- v <- type.try_make_transport_cps s0 ℤ;
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x4 <- (if (let (x4, _) := idc_args in x4) =? 1
- then
- Some
- (#(fancy_sell)%expr @
- (v (Compile.reflect x3), x0, x1))%expr_pat
- else None);
- Some (Base x4)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((s0 -> (projT1 args)) -> ℤ) -> ℤ)%ptype
+ then
+ v <- type.try_make_transport_cps s0 ℤ;
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x4 <- (if (let (x4, _) := idc_args in x4) =? 1
+ then
+ Some
+ (#(fancy_sell)%expr @
+ (v (Compile.reflect x3), x0, x1))%expr_pat
+ else None);
+ Some (Base x4)
+ else None
| None => None
end
| _ => None
@@ -1789,7 +2031,11 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((ℤ -> ℤ) -> ℤ)%ptype option (fun x2 : option => x2)
with
| Some (_, _, _) =>
- Some (Base (#(fancy_selc)%expr @ (x, x0, x1))%expr_pat)
+ if
+ type.type_beq base.type base.type.type_beq ((ℤ -> ℤ) -> ℤ)%ptype
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ then Some (Base (#(fancy_selc)%expr @ (x, x0, x1))%expr_pat)
+ else None
| None => None
end);;
None);;;
@@ -1801,7 +2047,11 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((ℤ -> ℤ) -> ℤ)%ptype option (fun x2 : option => x2)
with
| Some (_, _, _) =>
- Some (Base (#(fancy_addm)%expr @ (x, x0, x1))%expr_pat)
+ if
+ type.type_beq base.type base.type.type_beq ((ℤ -> ℤ) -> ℤ)%ptype
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ then Some (Base (#(fancy_addm)%expr @ (x, x0, x1))%expr_pat)
+ else None
| None => None
end;;;
Base (#(Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat)%option
@@ -1819,21 +2069,27 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- x3 <- (if
- (let (x3, _) := idc_args in x3) =?
- 2 ^ Z.log2 (let (x3, _) := idc_args in x3)
- then
- Some
- (#(fancy_rshi
- (Z.log2 (let (x3, _) := idc_args in x3))
- (let (x3, _) := idc_args0 in x3))%expr @
- (x0, x1))%expr_pat
- else None);
- Some (Base x3)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args0) -> ℤ) -> ℤ) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ x3 <- (if
+ (let (x3, _) := idc_args in x3) =?
+ 2 ^ Z.log2 (let (x3, _) := idc_args in x3)
+ then
+ Some
+ (#(fancy_rshi
+ (Z.log2 (let (x3, _) := idc_args in x3))
+ (let (x3, _) := idc_args0 in x3))%expr @
+ (x0, x1))%expr_pat
+ else None);
+ Some (Base x3)
+ else None
| None => None
end
| _ => None
diff --git a/src/Experiments/NewPipeline/nbe_rewrite_head.out b/src/Experiments/NewPipeline/nbe_rewrite_head.out
index ae46311be..f56e12da0 100644
--- a/src/Experiments/NewPipeline/nbe_rewrite_head.out
+++ b/src/Experiments/NewPipeline/nbe_rewrite_head.out
@@ -11,8 +11,12 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x0 : option => x0)
with
| Some _ =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- Some (Base (##(Nat.succ (let (x0, _) := idc_args in x0)))%expr)
+ if type.type_beq base.type base.type.type_beq ℕ (projT1 args)
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ Some
+ (Base (##(Nat.succ (let (x0, _) := idc_args in x0)))%expr)
+ else None
| None => None
end
| _ => None
@@ -29,8 +33,12 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x0 : option => x0)
with
| Some _ =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- Some (Base (##(Nat.pred (let (x0, _) := idc_args in x0)))%expr)
+ if type.type_beq base.type base.type.type_beq ℕ (projT1 args)
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ Some
+ (Base (##(Nat.pred (let (x0, _) := idc_args in x0)))%expr)
+ else None
| None => None
end
| _ => None
@@ -51,14 +59,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##(Nat.max (let (x1, _) := idc_args in x1)
- (let (x1, _) := idc_args0 in x1)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℕ -> ℕ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##(Nat.max (let (x1, _) := idc_args in x1)
+ (let (x1, _) := idc_args0 in x1)))%expr)
+ else None
| None => None
end
| _ => None
@@ -81,14 +94,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##((let (x1, _) := idc_args in x1) *
- (let (x1, _) := idc_args0 in x1))%nat)%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℕ -> ℕ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##((let (x1, _) := idc_args in x1) *
+ (let (x1, _) := idc_args0 in x1))%nat)%expr)
+ else None
| None => None
end
| _ => None
@@ -111,14 +129,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##((let (x1, _) := idc_args in x1) +
- (let (x1, _) := idc_args0 in x1))%nat)%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℕ -> ℕ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##((let (x1, _) := idc_args in x1) +
+ (let (x1, _) := idc_args0 in x1))%nat)%expr)
+ else None
| None => None
end
| _ => None
@@ -141,14 +164,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##((let (x1, _) := idc_args in x1) -
- (let (x1, _) := idc_args0 in x1))%nat)%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℕ -> ℕ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##((let (x1, _) := idc_args in x1) -
+ (let (x1, _) := idc_args0 in x1))%nat)%expr)
+ else None
| None => None
end
| _ => None
@@ -178,14 +206,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
s0) -> s)%ptype option (fun x2 : option => x2)
with
| Some (_, _, _, (_, (_, (_, _)), b3, b2)) =>
- _ <- ident.unify pattern.ident.fst fst;
- _ <- ident.unify pattern.ident.pair pair;
- v <- type.try_make_transport_cps s0 b3;
- _ <- type.try_make_transport_cps s b2;
- v1 <- base.try_make_transport_cps b3 A;
- v2 <- base.try_make_transport_cps A A;
- v3 <- base.try_make_transport_cps A A;
- Some (Base (v3 (v2 (v1 (v (Compile.reflect x1))))))
+ if
+ type.type_beq base.type base.type.type_beq
+ (((b3 * b2)%etype -> b3) ->
+ ((b3 -> b2 -> (b3 * b2)%etype) -> b3) -> b2)%ptype
+ (((A * B)%etype -> A) ->
+ (((let (x2, _) := args in x2) ->
+ (let (_, y) := args in y) ->
+ ((let (x2, _) := args in x2) * (let (_, y) := args in y))%etype) ->
+ s0) -> s)%ptype
+ then
+ _ <- ident.unify pattern.ident.fst fst;
+ _ <- ident.unify pattern.ident.pair pair;
+ v <- type.try_make_transport_cps s0 b3;
+ _ <- type.try_make_transport_cps s b2;
+ v1 <- base.try_make_transport_cps b3 A;
+ v2 <- base.try_make_transport_cps A A;
+ v3 <- base.try_make_transport_cps A A;
+ Some (Base (v3 (v2 (v1 (v (Compile.reflect x1))))))
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
@@ -218,14 +257,25 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
s0) -> s)%ptype option (fun x2 : option => x2)
with
| Some (_, _, _, (_, (_, (_, _)), b3, b2)) =>
- _ <- ident.unify pattern.ident.snd snd;
- _ <- ident.unify pattern.ident.pair pair;
- _ <- type.try_make_transport_cps s0 b3;
- v0 <- type.try_make_transport_cps s b2;
- v1 <- base.try_make_transport_cps b2 B;
- v2 <- base.try_make_transport_cps B B;
- v3 <- base.try_make_transport_cps B B;
- Some (Base (v3 (v2 (v1 (v0 (Compile.reflect x0))))))
+ if
+ type.type_beq base.type base.type.type_beq
+ (((b3 * b2)%etype -> b2) ->
+ ((b3 -> b2 -> (b3 * b2)%etype) -> b3) -> b2)%ptype
+ (((A * B)%etype -> B) ->
+ (((let (x2, _) := args in x2) ->
+ (let (_, y) := args in y) ->
+ ((let (x2, _) := args in x2) * (let (_, y) := args in y))%etype) ->
+ s0) -> s)%ptype
+ then
+ _ <- ident.unify pattern.ident.snd snd;
+ _ <- ident.unify pattern.ident.pair pair;
+ _ <- type.try_make_transport_cps s0 b3;
+ v0 <- type.try_make_transport_cps s b2;
+ v1 <- base.try_make_transport_cps b2 B;
+ v2 <- base.try_make_transport_cps B B;
+ v3 <- base.try_make_transport_cps B B;
+ Some (Base (v3 (v2 (v1 (v0 (Compile.reflect x0))))))
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
@@ -262,25 +312,38 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
with
| Some
(_, (_, _), (_, _, _), (_, (_, b7)), (_, (_, (_, _)), b9, b8)) =>
- _ <- ident.unify pattern.ident.prod_rect prod_rect;
- base.try_make_transport_cps A b9
- (fun a13 : option =>
- (fa <- (fun (T0 : Type) (k : option -> T0) =>
- match a13 with
- | Some x' =>
- base.try_make_transport_cps B b8
- (fun a14 : option =>
- fa <- (fun (T1 : Type) (k0 : option -> T1) =>
- match a14 with
- | Some x'0 =>
- base.try_make_transport_cps T b7
- (fun a15 : option =>
- fa <- (fun (T2 : Type)
- (k1 : option -> T2) =>
- match a15 with
- | Some x'1 =>
- (return Some
- (fun
+ if
+ type.type_beq base.type base.type.type_beq
+ ((((b9 -> b8 -> b7) -> (b9 * b8)%etype -> b7) ->
+ b9 -> b8 -> b7) ->
+ ((b9 -> b8 -> (b9 * b8)%etype) -> b9) -> b8)%ptype
+ ((((A -> B -> T) -> (A * B)%etype -> T) -> A -> B -> T) ->
+ (((let (x3, _) := args in x3) ->
+ (let (_, y) := args in y) ->
+ ((let (x3, _) := args in x3) * (let (_, y) := args in y))%etype) ->
+ s0) -> s)%ptype
+ then
+ _ <- ident.unify pattern.ident.prod_rect prod_rect;
+ base.try_make_transport_cps A b9
+ (fun a13 : option =>
+ (fa <- (fun (T0 : Type) (k : option -> T0) =>
+ match a13 with
+ | Some x' =>
+ base.try_make_transport_cps B b8
+ (fun a14 : option =>
+ fa <- (fun (T1 : Type) (k0 : option -> T1)
+ =>
+ match a14 with
+ | Some x'0 =>
+ base.try_make_transport_cps T b7
+ (fun a15 : option =>
+ fa <- (fun (T2 : Type)
+ (k1 : option -> T2)
+ =>
+ match a15 with
+ | Some x'1 =>
+ (return Some
+ (fun
v :
expr b9 ->
expr B ->
@@ -289,43 +352,44 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
=>
x'1
(x'0 v)))
- T2 k1
- | None => k1 None
- end);
- k0 fa)
- | None => k0 None
- end);
- fa0 <- (fun (T1 : Type) (k0 : option -> T1)
- =>
- match fa with
- | Some x'0 =>
- (return Some
- (fun
- v : expr A ->
- expr B ->
- UnderLets
- (expr T) =>
- x'0 (x' v))) T1 k0
+ T2 k1
+ | None => k1 None
+ end);
+ k0 fa)
| None => k0 None
end);
- k fa0)
- | None => k None
- end);
- match fa with
- | Some v =>
- (_ <- ident.unify pattern.ident.pair pair;
- v0 <- type.try_make_transport_cps s0 b9;
- v1 <- type.try_make_transport_cps s b8;
- v2 <- base.try_make_transport_cps b7 T;
- v3 <- base.try_make_transport_cps T T;
- v4 <- base.try_make_transport_cps T T;
- Some
- (fv <-- v2
- (v x (v0 (Compile.reflect x2))
- (v1 (Compile.reflect x1)));
- Base (v4 (v3 fv)))%under_lets)%option
- | None => None
- end)%cps)
+ fa0 <- (fun (T1 : Type) (k0 : option -> T1)
+ =>
+ match fa with
+ | Some x'0 =>
+ (return Some
+ (fun
+ v : expr A ->
+ expr B ->
+ UnderLets
+ (expr T) =>
+ x'0 (x' v))) T1 k0
+ | None => k0 None
+ end);
+ k fa0)
+ | None => k None
+ end);
+ match fa with
+ | Some v =>
+ (_ <- ident.unify pattern.ident.pair pair;
+ v0 <- type.try_make_transport_cps s0 b9;
+ v1 <- type.try_make_transport_cps s b8;
+ v2 <- base.try_make_transport_cps b7 T;
+ v3 <- base.try_make_transport_cps T T;
+ v4 <- base.try_make_transport_cps T T;
+ Some
+ (fv <-- v2
+ (v x (v0 (Compile.reflect x2))
+ (v1 (Compile.reflect x1)));
+ Base (v4 (v3 fv)))%under_lets)%option
+ | None => None
+ end)%cps)
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
@@ -358,49 +422,58 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, (_, _, (_, _)), (_, _), (_, b8), _) =>
- _ <- ident.unify pattern.ident.bool_rect bool_rect;
- base.try_make_transport_cps T b8
- (fun a9 : option =>
- (fa <- (fun (T0 : Type) (k : option -> T0) =>
- match a9 with
- | Some x' =>
- (return Some
- (fun
- v : expr unit -> UnderLets (expr T)
- => x' v)) T0 k
- | None => k None
- end);
- match fa with
- | Some v =>
- base.try_make_transport_cps T b8
- (fun a10 : option =>
- fa0 <- (fun (T0 : Type) (k : option -> T0) =>
- match a10 with
- | Some x' =>
- (return Some
- (fun
- v0 : expr unit ->
- UnderLets (expr T) =>
- x' v0)) T0 k
- | None => k None
- end);
- match fa0 with
- | Some v0 =>
- (idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v1 <- base.try_make_transport_cps b8 T;
- v2 <- base.try_make_transport_cps T T;
- v3 <- base.try_make_transport_cps T T;
- Some
- (fv <-- v1
- (if let (x2, _) := idc_args0 in x2
- then v x (##tt)%expr
- else v0 x0 (##tt)%expr);
- Base (v3 (v2 fv)))%under_lets)%option
- | None => None
- end)
- | None => None
- end)%cps)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((((unit -> b8) -> (unit -> b8) -> bool -> b8) ->
+ unit -> b8) -> unit -> b8) -> bool)%ptype
+ (((((unit -> T) -> (unit -> T) -> bool -> T) -> unit -> T) ->
+ unit -> T) -> (projT1 args))%ptype
+ then
+ _ <- ident.unify pattern.ident.bool_rect bool_rect;
+ base.try_make_transport_cps T b8
+ (fun a9 : option =>
+ (fa <- (fun (T0 : Type) (k : option -> T0) =>
+ match a9 with
+ | Some x' =>
+ (return Some
+ (fun
+ v : expr unit ->
+ UnderLets (expr T) =>
+ x' v)) T0 k
+ | None => k None
+ end);
+ match fa with
+ | Some v =>
+ base.try_make_transport_cps T b8
+ (fun a10 : option =>
+ fa0 <- (fun (T0 : Type) (k : option -> T0) =>
+ match a10 with
+ | Some x' =>
+ (return Some
+ (fun
+ v0 : expr unit ->
+ UnderLets (expr T) =>
+ x' v0)) T0 k
+ | None => k None
+ end);
+ match fa0 with
+ | Some v0 =>
+ (idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v1 <- base.try_make_transport_cps b8 T;
+ v2 <- base.try_make_transport_cps T T;
+ v3 <- base.try_make_transport_cps T T;
+ Some
+ (fv <-- v1
+ (if let (x2, _) := idc_args0 in x2
+ then v x (##tt)%expr
+ else v0 x0 (##tt)%expr);
+ Base (v3 (v2 fv)))%under_lets)%option
+ | None => None
+ end)
+ | None => None
+ end)%cps)
+ else None
| None => None
end
| _ => None
@@ -428,38 +501,194 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x2 : option => x2)
with
| Some (_, _, (_, (_, _), (_, _)), (_, _), (_, (_, b10)), _) =>
- _ <- ident.unify pattern.ident.nat_rect nat_rect;
- base.try_make_transport_cps P b10
- (fun a11 : option =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (((((unit -> b10) -> (ℕ -> b10 -> b10) -> ℕ -> b10) ->
+ unit -> b10) -> ℕ -> b10 -> b10) -> ℕ)%ptype
+ (((((unit -> P) -> (ℕ -> P -> P) -> ℕ -> P) -> unit -> P) ->
+ ℕ -> P -> P) -> (projT1 args))%ptype
+ then
+ _ <- ident.unify pattern.ident.nat_rect nat_rect;
+ base.try_make_transport_cps P b10
+ (fun a11 : option =>
+ (fa <- (fun (T : Type) (k : option -> T) =>
+ match a11 with
+ | Some x' =>
+ (return Some
+ (fun
+ v : expr unit ->
+ UnderLets (expr P) =>
+ x' v)) T k
+ | None => k None
+ end);
+ match fa with
+ | Some v =>
+ base.try_make_transport_cps P b10
+ (fun a12 : option =>
+ fa0 <- (fun (T : Type) (k : option -> T) =>
+ match a12 with
+ | Some x' =>
+ base.try_make_transport_cps P b10
+ (fun a13 : option =>
+ fa0 <- (fun (T0 : Type)
+ (k0 : option -> T0) =>
+ match a13 with
+ | Some x'0 =>
+ (return Some
+ (fun
+ v0 :
+ expr ℕ ->
+ expr P ->
+ UnderLets
+ (expr P)
+ =>
+ x'0 (x' v0)))
+ T0 k0
+ | None => k0 None
+ end);
+ k fa0)
+ | None => k None
+ end);
+ fa1 <- (fun (T : Type) (k : option -> T) =>
+ match fa0 with
+ | Some x' =>
+ (return Some
+ (fun
+ v0 : expr ℕ ->
+ expr P ->
+ UnderLets (expr P) =>
+ x' v0)) T k
+ | None => k None
+ end);
+ match fa1 with
+ | Some v0 =>
+ (idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v1 <- base.try_make_transport_cps b10 P;
+ v2 <- base.try_make_transport_cps P P;
+ v3 <- base.try_make_transport_cps P P;
+ Some
+ (fv <-- v1
+ (Datatypes.nat_rect
+ (fun _ : nat =>
+ UnderLets
+ (expr
+ (base.subst_default
+ '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b10
+ (PositiveMap.empty
+ base.type)))))
+ (v x (##tt)%expr)
+ (fun (n' : nat)
+ (rec : UnderLets
+ (expr
+ (base.subst_default
+ '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive)
+ b10
+ (PositiveMap.empty
+ base.type)))))
+ =>
+ rec0 <-- rec;
+ v0 x0 (##n')%expr rec0)
+ (let (x2, _) := idc_args0 in x2));
+ Base (v3 (v2 fv)))%under_lets)%option
+ | None => None
+ end)
+ | None => None
+ end)%cps)
+ else None
+ | None => None
+ end
+ | _ => None
+ end;;
+ None);;;
+ Base
+ (#(nat_rect)%expr @ (λ x2 : var unit,
+ to_expr (x ($x2)))%expr @
+ (λ (x2 : var ℕ)(x3 : var P),
+ to_expr (x0 ($x2) ($x3)))%expr @ x1)%expr_pat)%option
+| @nat_rect_arrow P Q =>
+ fun (x : expr P -> UnderLets (expr Q))
+ (x0 : expr ℕ ->
+ (expr P -> UnderLets (expr Q)) -> expr P -> UnderLets (expr Q))
+ (x1 : expr ℕ) (x2 : expr P) =>
+ (match x1 with
+ | @expr.Ident _ _ _ t idc =>
+ args <- invert_bind_args idc Raw.ident.Literal;
+ match
+ pattern.type.unify_extracted_cps
+ (((((('1%pbtype -> '2%pbtype) ->
+ (ℕ -> ('1%pbtype -> '2%pbtype) -> '1%pbtype -> '2%pbtype) ->
+ ℕ -> '1%pbtype -> '2%pbtype) -> '1%pbtype -> '2%pbtype) ->
+ ℕ -> ('1%pbtype -> '2%pbtype) -> '1%pbtype -> '2%pbtype) ->
+ ℕ) -> '1%pbtype)%ptype
+ ((((((P -> Q) -> (ℕ -> (P -> Q) -> P -> Q) -> ℕ -> P -> Q) ->
+ P -> Q) -> ℕ -> (P -> Q) -> P -> Q) -> (projT1 args)) ->
+ P)%ptype option (fun x3 : option => x3)
+ with
+ | Some
+ (_, _, (_, (_, _, (_, _)), (_, (_, _))), (_, _),
+ (_, (_, _, (_, b16))), _, b) =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((((((b -> b16) ->
+ (ℕ -> (b -> b16) -> b -> b16) -> ℕ -> b -> b16) ->
+ b -> b16) -> ℕ -> (b -> b16) -> b -> b16) -> ℕ) ->
+ b)%ptype
+ ((((((P -> Q) -> (ℕ -> (P -> Q) -> P -> Q) -> ℕ -> P -> Q) ->
+ P -> Q) -> ℕ -> (P -> Q) -> P -> Q) -> (projT1 args)) ->
+ P)%ptype
+ then
+ _ <- ident.unify pattern.ident.nat_rect_arrow nat_rect_arrow;
+ base.try_make_transport_cps P b
+ (fun a17 : option =>
(fa <- (fun (T : Type) (k : option -> T) =>
- match a11 with
+ match a17 with
| Some x' =>
- (return Some
- (fun
- v : expr unit -> UnderLets (expr P)
- => x' v)) T k
+ base.try_make_transport_cps Q b16
+ (fun a18 : option =>
+ fa <- (fun (T0 : Type) (k0 : option -> T0) =>
+ match a18 with
+ | Some x'0 =>
+ (return Some
+ (fun
+ v : expr P ->
+ UnderLets
+ (expr Q) =>
+ x'0 (x' v))) T0 k0
+ | None => k0 None
+ end);
+ k fa)
| None => k None
end);
match fa with
| Some v =>
- base.try_make_transport_cps P b10
- (fun a12 : option =>
+ base.try_make_transport_cps P b
+ (fun a18 : option =>
fa0 <- (fun (T : Type) (k : option -> T) =>
- match a12 with
+ match a18 with
| Some x' =>
- base.try_make_transport_cps P b10
- (fun a13 : option =>
+ base.try_make_transport_cps Q b16
+ (fun a19 : option =>
fa0 <- (fun (T0 : Type)
(k0 : option -> T0) =>
- match a13 with
+ match a19 with
| Some x'0 =>
(return Some
(fun
v0 :
expr ℕ ->
+ (expr P ->
+ UnderLets
+ (expr Q)) ->
expr P ->
UnderLets
- (expr P)
+ (expr Q)
=> x'0 (x' v0)))
T0 k0
| None => k0 None
@@ -470,248 +699,141 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
fa1 <- (fun (T : Type) (k : option -> T) =>
match fa0 with
| Some x' =>
+ base.try_make_transport_cps P b
+ (fun a19 : option =>
+ fa1 <- (fun (T0 : Type)
+ (k0 : option -> T0) =>
+ match a19 with
+ | Some x'0 =>
+ base.try_make_transport_cps
+ Q b16
+ (fun a20 : option =>
+ fa1 <- (fun
+ (T1 : Type)
+ (k1 :
+ option ->
+ T1) =>
+ match a20 with
+ | Some x'1 =>
+ (return
+ Some
+ (fun
+ v0 :
+ expr ℕ ->
+ (expr b ->
+ UnderLets
+ (expr b16)) ->
+ expr P ->
+ UnderLets
+ (expr Q)
+ =>
+ x'1
+ (x'0 v0)))
+ T1 k1
+ | None =>
+ k1 None
+ end);
+ k0 fa1)
+ | None => k0 None
+ end);
+ fa2 <- (fun (T0 : Type)
+ (k0 : option -> T0) =>
+ match fa1 with
+ | Some x'0 =>
+ (return Some
+ (fun
+ v0 :
+ expr ℕ ->
+ (expr P ->
+ UnderLets
+ (expr Q)) ->
+ expr P ->
+ UnderLets
+ (expr Q)
+ => x'0 (x' v0)))
+ T0 k0
+ | None => k0 None
+ end);
+ k fa2)
+ | None => k None
+ end);
+ fa2 <- (fun (T : Type) (k : option -> T) =>
+ match fa1 with
+ | Some x' =>
(return Some
(fun
v0 : expr ℕ ->
+ (expr P ->
+ UnderLets (expr Q)) ->
expr P ->
- UnderLets (expr P) =>
+ UnderLets (expr Q) =>
x' v0)) T k
| None => k None
end);
- match fa1 with
+ match fa2 with
| Some v0 =>
(idc_args0 <- ident.unify pattern.ident.Literal
##(projT2 args);
- v1 <- base.try_make_transport_cps b10 P;
- v2 <- base.try_make_transport_cps P P;
- v3 <- base.try_make_transport_cps P P;
+ v1 <- base.try_make_transport_cps P b;
+ v2 <- base.try_make_transport_cps b16 Q;
+ v3 <- base.try_make_transport_cps Q Q;
+ v4 <- base.try_make_transport_cps Q Q;
Some
- (fv <-- v1
+ (fv <-- v2
(Datatypes.nat_rect
(fun _ : nat =>
+ expr
+ (base.subst_default '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 2%positive) b16
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b
+ (PositiveMap.empty
+ base.type)))) ->
UnderLets
(expr
(base.subst_default
- '1
+ '2
(PositiveMap.add
(PositiveSet.rev
- 1%positive) b10
- (PositiveMap.empty
- base.type)))))
- (v x (##tt)%expr)
+ 2%positive) b16
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b
+ (PositiveMap.empty
+ base.type))))))
+ (v x)
(fun (n' : nat)
- (rec : UnderLets
+ (rec : expr
+ (base.subst_default
+ '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 2%positive)
+ b16
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive)
+ b
+ (PositiveMap.empty
+ base.type)))) ->
+ UnderLets
(expr
(base.subst_default
- '1
+ '2
(PositiveMap.add
(PositiveSet.rev
+ 2%positive)
+ b16
+ (PositiveMap.add
+ (PositiveSet.rev
1%positive)
- b10
- (PositiveMap.empty
- base.type)))))
- =>
- rec0 <-- rec;
- v0 x0 (##n')%expr rec0)
- (let (x2, _) := idc_args0 in x2));
- Base (v3 (v2 fv)))%under_lets)%option
- | None => None
- end)
- | None => None
- end)%cps)
- | None => None
- end
- | _ => None
- end;;
- None);;;
- Base
- (#(nat_rect)%expr @ (λ x2 : var unit,
- to_expr (x ($x2)))%expr @
- (λ (x2 : var ℕ)(x3 : var P),
- to_expr (x0 ($x2) ($x3)))%expr @ x1)%expr_pat)%option
-| @nat_rect_arrow P Q =>
- fun (x : expr P -> UnderLets (expr Q))
- (x0 : expr ℕ ->
- (expr P -> UnderLets (expr Q)) -> expr P -> UnderLets (expr Q))
- (x1 : expr ℕ) (x2 : expr P) =>
- (match x1 with
- | @expr.Ident _ _ _ t idc =>
- args <- invert_bind_args idc Raw.ident.Literal;
- match
- pattern.type.unify_extracted_cps
- (((((('1%pbtype -> '2%pbtype) ->
- (ℕ -> ('1%pbtype -> '2%pbtype) -> '1%pbtype -> '2%pbtype) ->
- ℕ -> '1%pbtype -> '2%pbtype) -> '1%pbtype -> '2%pbtype) ->
- ℕ -> ('1%pbtype -> '2%pbtype) -> '1%pbtype -> '2%pbtype) ->
- ℕ) -> '1%pbtype)%ptype
- ((((((P -> Q) -> (ℕ -> (P -> Q) -> P -> Q) -> ℕ -> P -> Q) ->
- P -> Q) -> ℕ -> (P -> Q) -> P -> Q) -> (projT1 args)) ->
- P)%ptype option (fun x3 : option => x3)
- with
- | Some
- (_, _, (_, (_, _, (_, _)), (_, (_, _))), (_, _),
- (_, (_, _, (_, b16))), _, b) =>
- _ <- ident.unify pattern.ident.nat_rect_arrow nat_rect_arrow;
- base.try_make_transport_cps P b
- (fun a17 : option =>
- (fa <- (fun (T : Type) (k : option -> T) =>
- match a17 with
- | Some x' =>
- base.try_make_transport_cps Q b16
- (fun a18 : option =>
- fa <- (fun (T0 : Type) (k0 : option -> T0) =>
- match a18 with
- | Some x'0 =>
- (return Some
- (fun
- v : expr P ->
- UnderLets (expr Q)
- => x'0 (x' v))) T0 k0
- | None => k0 None
- end);
- k fa)
- | None => k None
- end);
- match fa with
- | Some v =>
- base.try_make_transport_cps P b
- (fun a18 : option =>
- fa0 <- (fun (T : Type) (k : option -> T) =>
- match a18 with
- | Some x' =>
- base.try_make_transport_cps Q b16
- (fun a19 : option =>
- fa0 <- (fun (T0 : Type)
- (k0 : option -> T0) =>
- match a19 with
- | Some x'0 =>
- (return Some
- (fun
- v0 :
- expr ℕ ->
- (expr P ->
- UnderLets
- (expr Q)) ->
- expr P ->
- UnderLets
- (expr Q)
- => x'0 (x' v0)))
- T0 k0
- | None => k0 None
- end);
- k fa0)
- | None => k None
- end);
- fa1 <- (fun (T : Type) (k : option -> T) =>
- match fa0 with
- | Some x' =>
- base.try_make_transport_cps P b
- (fun a19 : option =>
- fa1 <- (fun (T0 : Type)
- (k0 : option -> T0) =>
- match a19 with
- | Some x'0 =>
- base.try_make_transport_cps
- Q b16
- (fun a20 : option =>
- fa1 <- (fun (T1 : Type)
- (k1 :
- option -> T1)
- =>
- match a20 with
- | Some x'1 =>
- (return
- Some
- (fun
- v0 :
- expr ℕ ->
- (expr b ->
- UnderLets
- (expr b16)) ->
- expr P ->
- UnderLets
- (expr Q)
- =>
- x'1
- (x'0 v0)))
- T1 k1
- | None =>
- k1 None
- end);
- k0 fa1)
- | None => k0 None
- end);
- fa2 <- (fun (T0 : Type)
- (k0 : option -> T0) =>
- match fa1 with
- | Some x'0 =>
- (return Some
- (fun
- v0 :
- expr ℕ ->
- (expr P ->
- UnderLets
- (expr Q)) ->
- expr P ->
- UnderLets
- (expr Q)
- => x'0 (x' v0)))
- T0 k0
- | None => k0 None
- end);
- k fa2)
- | None => k None
- end);
- fa2 <- (fun (T : Type) (k : option -> T) =>
- match fa1 with
- | Some x' =>
- (return Some
- (fun
- v0 : expr ℕ ->
- (expr P ->
- UnderLets (expr Q)) ->
- expr P ->
- UnderLets (expr Q) =>
- x' v0)) T k
- | None => k None
- end);
- match fa2 with
- | Some v0 =>
- (idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v1 <- base.try_make_transport_cps P b;
- v2 <- base.try_make_transport_cps b16 Q;
- v3 <- base.try_make_transport_cps Q Q;
- v4 <- base.try_make_transport_cps Q Q;
- Some
- (fv <-- v2
- (Datatypes.nat_rect
- (fun _ : nat =>
- expr
- (base.subst_default '1
- (PositiveMap.add
- (PositiveSet.rev
- 2%positive) b16
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type)))) ->
- UnderLets
- (expr
- (base.subst_default
- '2
- (PositiveMap.add
- (PositiveSet.rev
- 2%positive) b16
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type))))))
- (v x)
- (fun (n' : nat)
- (rec : expr
+ b
+ (PositiveMap.empty
+ base.type))))))
+ (v5 : expr
(base.subst_default
'1
(PositiveMap.add
@@ -723,42 +845,16 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
1%positive)
b
(PositiveMap.empty
- base.type)))) ->
- UnderLets
- (expr
- (base.subst_default
- '2
- (PositiveMap.add
- (PositiveSet.rev
- 2%positive)
- b16
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive)
- b
- (PositiveMap.empty
- base.type))))))
- (v5 : expr
- (base.subst_default
- '1
- (PositiveMap.add
- (PositiveSet.rev
- 2%positive)
- b16
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive)
- b
- (PositiveMap.empty
- base.type)))))
- => v0 x0 (##n')%expr rec v5)
- (let (x3, _) := idc_args0 in x3)
- (v1 x2));
- Base (v4 (v3 fv)))%under_lets)%option
- | None => None
- end)
- | None => None
- end)%cps)
+ base.type)))))
+ => v0 x0 (##n')%expr rec v5)
+ (let (x3, _) := idc_args0 in x3)
+ (v1 x2));
+ Base (v4 (v3 fv)))%under_lets)%option
+ | None => None
+ end)
+ | None => None
+ end)%cps)
+ else None
| None => None
end
| _ => None
@@ -789,53 +885,61 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
with
| Some
(_, _, (_, (_, (_, _)), (_, _)), (_, _), (_, (_, (_, b12))), b) =>
- _ <- ident.unify pattern.ident.list_rect list_rect;
- base.try_make_transport_cps P b12
- (fun a13 : option =>
- (fa <- (fun (T : Type) (k : option -> T) =>
- match a13 with
- | Some x' =>
- (return Some
- (fun v : expr unit -> UnderLets (expr P)
- => x' v)) T k
- | None => k None
- end);
- match fa with
- | Some v =>
- base.try_make_transport_cps A b
- (fun a14 : option =>
- fa0 <- (fun (T : Type) (k : option -> T) =>
- match a14 with
- | Some x' =>
- base.try_make_transport_cps A b
- (fun a15 : option =>
- fa0 <- (fun (T0 : Type)
- (k0 : option -> T0) =>
- match a15 with
- | Some x'0 =>
- base.try_make_transport_cps P
- b12
- (fun a16 : option =>
- fa0 <- (fun (T1 : Type)
- (k1 : option ->
- T1) =>
- match a16 with
- | Some x'1 =>
- base.try_make_transport_cps
- P b12
- (fun
- a17 : option
- =>
- fa0 <-
+ if
+ type.type_beq base.type base.type.type_beq
+ (((((unit -> b12) ->
+ (b -> (list b) -> b12 -> b12) -> (list b) -> b12) ->
+ unit -> b12) -> b -> (list b) -> b12 -> b12) -> (list b))%ptype
+ (((((unit -> P) -> (A -> (list A) -> P -> P) -> (list A) -> P) ->
+ unit -> P) -> A -> (list A) -> P -> P) -> (list A))%ptype
+ then
+ _ <- ident.unify pattern.ident.list_rect list_rect;
+ base.try_make_transport_cps P b12
+ (fun a13 : option =>
+ (fa <- (fun (T : Type) (k : option -> T) =>
+ match a13 with
+ | Some x' =>
+ (return Some
+ (fun v : expr unit -> UnderLets (expr P)
+ => x' v)) T k
+ | None => k None
+ end);
+ match fa with
+ | Some v =>
+ base.try_make_transport_cps A b
+ (fun a14 : option =>
+ fa0 <- (fun (T : Type) (k : option -> T) =>
+ match a14 with
+ | Some x' =>
+ base.try_make_transport_cps A b
+ (fun a15 : option =>
+ fa0 <- (fun (T0 : Type)
+ (k0 : option -> T0) =>
+ match a15 with
+ | Some x'0 =>
+ base.try_make_transport_cps
+ P b12
+ (fun a16 : option =>
+ fa0 <- (fun (T1 : Type)
+ (k1 : option ->
+ T1) =>
+ match a16 with
+ | Some x'1 =>
+ base.try_make_transport_cps
+ P b12
(fun
+ a17 : option
+ =>
+ fa0 <-
+ (fun
(T2 : Type)
(k2 :
option ->
T2) =>
- match
+ match
a17
- with
- | Some
+ with
+ | Some
x'2 =>
(return
Some
@@ -851,20 +955,20 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
x'2
(x'1 v0)))
T2 k2
- | None =>
+ | None =>
k2 None
- end);
- k1 fa0)
- | None => k1 None
- end);
- fa1 <- (fun (T1 : Type)
- (k1 : option ->
- T1) =>
- match fa0 with
- | Some x'1 =>
- (return
- Some
- (fun
+ end);
+ k1 fa0)
+ | None => k1 None
+ end);
+ fa1 <- (fun (T1 : Type)
+ (k1 : option ->
+ T1) =>
+ match fa0 with
+ | Some x'1 =>
+ (return
+ Some
+ (fun
v0 :
expr b ->
expr
@@ -872,45 +976,71 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
expr P ->
UnderLets
(expr P)
- =>
- x'1
+ =>
+ x'1
(x'0 v0)))
- T1 k1
- | None => k1 None
- end);
- k0 fa1)
- | None => k0 None
- end);
- fa1 <- (fun (T0 : Type)
- (k0 : option -> T0) =>
- match fa0 with
- | Some x'0 =>
- (return Some
- (fun
- v0 : expr A ->
- expr
+ T1 k1
+ | None => k1 None
+ end);
+ k0 fa1)
+ | None => k0 None
+ end);
+ fa1 <- (fun (T0 : Type)
+ (k0 : option -> T0) =>
+ match fa0 with
+ | Some x'0 =>
+ (return Some
+ (fun
+ v0 : expr A ->
+ expr
(list A) ->
- expr P ->
- UnderLets
+ expr P ->
+ UnderLets
(expr P)
- => x'0 (x' v0)))
- T0 k0
- | None => k0 None
- end);
- k fa1)
- | None => k None
- end);
- match fa0 with
- | Some v0 =>
- v1 <- base.try_make_transport_cps A b;
- v2 <- base.try_make_transport_cps b12 P;
- (fv <- v2
- (ls <- reflect_list (v1 x1);
- Some
- (Datatypes.list_rect
- (fun
- _ : Datatypes.list
- (expr
+ => x'0 (x' v0)))
+ T0 k0
+ | None => k0 None
+ end);
+ k fa1)
+ | None => k None
+ end);
+ match fa0 with
+ | Some v0 =>
+ v1 <- base.try_make_transport_cps A b;
+ v2 <- base.try_make_transport_cps b12 P;
+ (fv <- v2
+ (ls <- reflect_list (v1 x1);
+ Some
+ (Datatypes.list_rect
+ (fun
+ _ : Datatypes.list
+ (expr
+ (base.subst_default
+ '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 2%positive) b12
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive)
+ b
+ (PositiveMap.empty
+ base.type)))))
+ =>
+ UnderLets
+ (expr
+ (base.subst_default '2
+ (PositiveMap.add
+ (PositiveSet.rev
+ 2%positive) b12
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b
+ (PositiveMap.empty
+ base.type))))))
+ (v x (##tt)%expr)
+ (fun
+ (x2 : expr
(base.subst_default
'1
(PositiveMap.add
@@ -921,49 +1051,10 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
1%positive) b
(PositiveMap.empty
base.type)))))
- =>
- UnderLets
- (expr
- (base.subst_default '2
- (PositiveMap.add
- (PositiveSet.rev
- 2%positive) b12
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type))))))
- (v x (##tt)%expr)
- (fun
- (x2 : expr
- (base.subst_default
- '1
- (PositiveMap.add
- (PositiveSet.rev
- 2%positive) b12
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type)))))
- (xs : Datatypes.list
- (expr
- (base.subst_default
- '1
- (PositiveMap.add
- (PositiveSet.rev
- 2%positive)
- b12
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive)
- b
- (PositiveMap.empty
- base.type))))))
- (rec : UnderLets
+ (xs : Datatypes.list
(expr
(base.subst_default
- '2
+ '1
(PositiveMap.add
(PositiveSet.rev
2%positive)
@@ -974,18 +1065,33 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
b
(PositiveMap.empty
base.type))))))
- =>
- (rec' <-- rec;
- v0 x0 x2 (Compilers.reify_list xs)
- rec')%under_lets) ls));
- v3 <- base.try_make_transport_cps P P;
- v4 <- base.try_make_transport_cps P P;
- Some (fv0 <-- fv;
- Base (v4 (v3 fv0)))%under_lets)%option
- | None => None
- end)
- | None => None
- end)%cps)
+ (rec : UnderLets
+ (expr
+ (base.subst_default
+ '2
+ (PositiveMap.add
+ (PositiveSet.rev
+ 2%positive)
+ b12
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive)
+ b
+ (PositiveMap.empty
+ base.type))))))
+ =>
+ (rec' <-- rec;
+ v0 x0 x2 (Compilers.reify_list xs)
+ rec')%under_lets) ls));
+ v3 <- base.try_make_transport_cps P P;
+ v4 <- base.try_make_transport_cps P P;
+ Some (fv0 <-- fv;
+ Base (v4 (v3 fv0)))%under_lets)%option
+ | None => None
+ end)
+ | None => None
+ end)%cps)
+ else None
| None => None
end;;
None);;;
@@ -1014,42 +1120,53 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x2 : option => x2)
with
| Some (_, _, (_, (_, _), (_, _)), (_, _), (_, (_, b10)), b) =>
- _ <- ident.unify pattern.ident.list_case list_case;
- base.try_make_transport_cps P b10
- (fun a11 : option =>
- (fa <- (fun (T : Type) (k : option -> T) =>
- match a11 with
- | Some x' =>
- (return Some
- (fun
- v : expr unit -> UnderLets (expr P)
- => x' v)) T k
- | None => k None
- end);
- match fa with
- | Some v =>
- base.try_make_transport_cps A b
- (fun a12 : option =>
- fa0 <- (fun (T : Type) (k : option -> T) =>
- match a12 with
- | Some x' =>
- base.try_make_transport_cps A b
- (fun a13 : option =>
- fa0 <- (fun (T0 : Type)
- (k0 : option -> T0) =>
- match a13 with
- | Some x'0 =>
- base.try_make_transport_cps
- P b10
- (fun a14 : option =>
- fa0 <- (fun
- (T1 : Type)
- (k1 :
- option ->
- T1) =>
- match a14 with
- | Some x'1 =>
- (return
+ if
+ type.type_beq base.type base.type.type_beq
+ (((((unit -> b10) ->
+ (b -> (list b) -> b10) -> (list b) -> b10) ->
+ unit -> b10) -> b -> (list b) -> b10) -> (list b))%ptype
+ (((((unit -> P) -> (A -> (list A) -> P) -> (list A) -> P) ->
+ unit -> P) -> A -> (list A) -> P) -> (list args))%ptype
+ then
+ _ <- ident.unify pattern.ident.list_case list_case;
+ base.try_make_transport_cps P b10
+ (fun a11 : option =>
+ (fa <- (fun (T : Type) (k : option -> T) =>
+ match a11 with
+ | Some x' =>
+ (return Some
+ (fun
+ v : expr unit ->
+ UnderLets (expr P) =>
+ x' v)) T k
+ | None => k None
+ end);
+ match fa with
+ | Some v =>
+ base.try_make_transport_cps A b
+ (fun a12 : option =>
+ fa0 <- (fun (T : Type) (k : option -> T) =>
+ match a12 with
+ | Some x' =>
+ base.try_make_transport_cps A b
+ (fun a13 : option =>
+ fa0 <- (fun (T0 : Type)
+ (k0 : option -> T0) =>
+ match a13 with
+ | Some x'0 =>
+ base.try_make_transport_cps
+ P b10
+ (fun a14 : option =>
+ fa0 <- (fun
+ (T1 : Type)
+ (k1 :
+ option ->
+ T1) =>
+ match
+ a14
+ with
+ | Some x'1 =>
+ (return
Some
(fun
v0 :
@@ -1062,44 +1179,46 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
x'1
(x'0 v0)))
T1 k1
- | None =>
- k1 None
- end);
- k0 fa0)
- | None => k0 None
- end);
- fa1 <- (fun (T0 : Type)
- (k0 : option -> T0) =>
- match fa0 with
- | Some x'0 =>
- (return Some
- (fun
- v0 :
- expr A ->
- expr
+ | None =>
+ k1 None
+ end);
+ k0 fa0)
+ | None => k0 None
+ end);
+ fa1 <- (fun (T0 : Type)
+ (k0 : option -> T0) =>
+ match fa0 with
+ | Some x'0 =>
+ (return Some
+ (fun
+ v0 :
+ expr A ->
+ expr
(list A) ->
- UnderLets
+ UnderLets
(expr P)
- => x'0 (x' v0)))
- T0 k0
- | None => k0 None
- end);
- k fa1)
- | None => k None
- end);
- match fa0 with
- | Some _ =>
- (_ <- ident.unify pattern.ident.nil nil;
- v1 <- base.try_make_transport_cps b10 P;
- v2 <- base.try_make_transport_cps P P;
- v3 <- base.try_make_transport_cps P P;
- Some
- (fv <-- v1 (v x (##tt)%expr);
- Base (v3 (v2 fv)))%under_lets)%option
- | None => None
- end)
- | None => None
- end)%cps)
+ =>
+ x'0 (x' v0)))
+ T0 k0
+ | None => k0 None
+ end);
+ k fa1)
+ | None => k None
+ end);
+ match fa0 with
+ | Some _ =>
+ (_ <- ident.unify pattern.ident.nil nil;
+ v1 <- base.try_make_transport_cps b10 P;
+ v2 <- base.try_make_transport_cps P P;
+ v3 <- base.try_make_transport_cps P P;
+ Some
+ (fv <-- v1 (v x (##tt)%expr);
+ Base (v3 (v2 fv)))%under_lets)%option
+ | None => None
+ end)
+ | None => None
+ end)%cps)
+ else None
| None => None
end
| @expr.App _ _ _ s _
@@ -1123,42 +1242,55 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
| Some
(_, _, (_, (_, _), (_, _)), (_, _), (_, (_, b10)),
(_, (_, _), _, b11)) =>
- _ <- ident.unify pattern.ident.list_case list_case;
- base.try_make_transport_cps P b10
- (fun a15 : option =>
- (fa <- (fun (T : Type) (k : option -> T) =>
- match a15 with
- | Some x' =>
- (return Some
- (fun
- v : expr unit -> UnderLets (expr P)
- => x' v)) T k
- | None => k None
- end);
- match fa with
- | Some _ =>
- base.try_make_transport_cps A b11
- (fun a16 : option =>
- fa0 <- (fun (T : Type) (k : option -> T) =>
- match a16 with
- | Some x' =>
- base.try_make_transport_cps A b11
- (fun a17 : option =>
- fa0 <- (fun (T0 : Type)
- (k0 : option -> T0) =>
- match a17 with
- | Some x'0 =>
- base.try_make_transport_cps
- P b10
- (fun a18 : option =>
- fa0 <- (fun
- (T1 : Type)
- (k1 :
- option ->
- T1) =>
- match a18 with
- | Some x'1 =>
- (return
+ if
+ type.type_beq base.type base.type.type_beq
+ (((((unit -> b10) ->
+ (b11 -> (list b11) -> b10) -> (list b11) -> b10) ->
+ unit -> b10) -> b11 -> (list b11) -> b10) ->
+ ((b11 -> (list b11) -> (list b11)) -> b11) -> (list b11))%ptype
+ (((((unit -> P) -> (A -> (list A) -> P) -> (list A) -> P) ->
+ unit -> P) -> A -> (list A) -> P) ->
+ ((args -> (list args) -> (list args)) -> s0) -> s)%ptype
+ then
+ _ <- ident.unify pattern.ident.list_case list_case;
+ base.try_make_transport_cps P b10
+ (fun a15 : option =>
+ (fa <- (fun (T : Type) (k : option -> T) =>
+ match a15 with
+ | Some x' =>
+ (return Some
+ (fun
+ v : expr unit ->
+ UnderLets (expr P) =>
+ x' v)) T k
+ | None => k None
+ end);
+ match fa with
+ | Some _ =>
+ base.try_make_transport_cps A b11
+ (fun a16 : option =>
+ fa0 <- (fun (T : Type) (k : option -> T) =>
+ match a16 with
+ | Some x' =>
+ base.try_make_transport_cps A b11
+ (fun a17 : option =>
+ fa0 <- (fun (T0 : Type)
+ (k0 : option -> T0) =>
+ match a17 with
+ | Some x'0 =>
+ base.try_make_transport_cps
+ P b10
+ (fun a18 : option =>
+ fa0 <- (fun
+ (T1 : Type)
+ (k1 :
+ option ->
+ T1) =>
+ match
+ a18
+ with
+ | Some x'1 =>
+ (return
Some
(fun
v0 :
@@ -1171,48 +1303,50 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
x'1
(x'0 v0)))
T1 k1
- | None =>
- k1 None
- end);
- k0 fa0)
- | None => k0 None
- end);
- fa1 <- (fun (T0 : Type)
- (k0 : option -> T0) =>
- match fa0 with
- | Some x'0 =>
- (return Some
- (fun
- v0 :
- expr A ->
- expr
+ | None =>
+ k1 None
+ end);
+ k0 fa0)
+ | None => k0 None
+ end);
+ fa1 <- (fun (T0 : Type)
+ (k0 : option -> T0) =>
+ match fa0 with
+ | Some x'0 =>
+ (return Some
+ (fun
+ v0 :
+ expr A ->
+ expr
(list A) ->
- UnderLets
+ UnderLets
(expr P)
- => x'0 (x' v0)))
- T0 k0
- | None => k0 None
- end);
- k fa1)
- | None => k None
- end);
- match fa0 with
- | Some v0 =>
- (_ <- ident.unify pattern.ident.cons cons;
- v1 <- type.try_make_transport_cps s0 b11;
- v2 <- type.try_make_transport_cps s (list b11);
- v3 <- base.try_make_transport_cps b10 P;
- v4 <- base.try_make_transport_cps P P;
- v5 <- base.try_make_transport_cps P P;
- Some
- (fv <-- v3
- (v0 x0 (v1 (Compile.reflect x3))
- (v2 (Compile.reflect x2)));
- Base (v5 (v4 fv)))%under_lets)%option
- | None => None
- end)
- | None => None
- end)%cps)
+ =>
+ x'0 (x' v0)))
+ T0 k0
+ | None => k0 None
+ end);
+ k fa1)
+ | None => k None
+ end);
+ match fa0 with
+ | Some v0 =>
+ (_ <- ident.unify pattern.ident.cons cons;
+ v1 <- type.try_make_transport_cps s0 b11;
+ v2 <- type.try_make_transport_cps s (list b11);
+ v3 <- base.try_make_transport_cps b10 P;
+ v4 <- base.try_make_transport_cps P P;
+ v5 <- base.try_make_transport_cps P P;
+ Some
+ (fv <-- v3
+ (v0 x0 (v1 (Compile.reflect x3))
+ (v2 (Compile.reflect x2)));
+ Base (v5 (v4 fv)))%under_lets)%option
+ | None => None
+ end)
+ | None => None
+ end)%cps)
+ else None
| None => None
end
| @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
@@ -1239,11 +1373,17 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(((list T) -> ℕ) -> (list T))%ptype option (fun x0 : option => x0)
with
| Some (_, _, b) =>
- _ <- ident.unify pattern.ident.List_length List_length;
- v <- base.try_make_transport_cps T b;
- x0 <- (xs <- reflect_list (v x);
- Some (##(length xs))%expr);
- Some (Base x0)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((list b) -> ℕ) -> (list b))%ptype
+ (((list T) -> ℕ) -> (list T))%ptype
+ then
+ _ <- ident.unify pattern.ident.List_length List_length;
+ v <- base.try_make_transport_cps T b;
+ x0 <- (xs <- reflect_list (v x);
+ Some (##(length xs))%expr);
+ Some (Base x0)
+ else None
| None => None
end;;;
Base (#(List_length)%expr @ x)%expr_pat)%option
@@ -1261,16 +1401,21 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (Compilers.reify_list
- (map (fun v : nat => (##v)%expr)
- (seq (let (x1, _) := idc_args in x1)
- (let (x1, _) := idc_args0 in x1)))))
+ if
+ type.type_beq base.type base.type.type_beq (ℕ -> ℕ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (Compilers.reify_list
+ (map (fun v : nat => (##v)%expr)
+ (seq (let (x1, _) := idc_args in x1)
+ (let (x1, _) := idc_args0 in x1)))))
+ else None
| None => None
end
| _ => None
@@ -1293,20 +1438,26 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x1 : option => x1)
with
| Some (_, (_, _), _, b) =>
- _ <- ident.unify pattern.ident.List_firstn List_firstn;
- idc_args0 <- ident.unify pattern.ident.Literal ##(projT2 args);
- v <- base.try_make_transport_cps A b;
- v0 <- base.try_make_transport_cps b A;
- fv <- v0
- (xs <- reflect_list (v x0);
- Some
- (Base
- (Compilers.reify_list
- (firstn (let (x1, _) := idc_args0 in x1) xs))));
- v1 <- base.try_make_transport_cps A A;
- v2 <- base.try_make_transport_cps A A;
- Some (fv0 <-- fv;
- Base (v2 (v1 fv0)))%under_lets
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℕ -> (list b) -> (list b)) -> ℕ) -> (list b))%ptype
+ (((ℕ -> (list A) -> (list A)) -> (projT1 args)) -> (list A))%ptype
+ then
+ _ <- ident.unify pattern.ident.List_firstn List_firstn;
+ idc_args0 <- ident.unify pattern.ident.Literal ##(projT2 args);
+ v <- base.try_make_transport_cps A b;
+ v0 <- base.try_make_transport_cps b A;
+ fv <- v0
+ (xs <- reflect_list (v x0);
+ Some
+ (Base
+ (Compilers.reify_list
+ (firstn (let (x1, _) := idc_args0 in x1) xs))));
+ v1 <- base.try_make_transport_cps A A;
+ v2 <- base.try_make_transport_cps A A;
+ Some (fv0 <-- fv;
+ Base (v2 (v1 fv0)))%under_lets
+ else None
| None => None
end
| _ => None
@@ -1327,20 +1478,26 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x1 : option => x1)
with
| Some (_, (_, _), _, b) =>
- _ <- ident.unify pattern.ident.List_skipn List_skipn;
- idc_args0 <- ident.unify pattern.ident.Literal ##(projT2 args);
- v <- base.try_make_transport_cps A b;
- v0 <- base.try_make_transport_cps b A;
- fv <- v0
- (xs <- reflect_list (v x0);
- Some
- (Base
- (Compilers.reify_list
- (skipn (let (x1, _) := idc_args0 in x1) xs))));
- v1 <- base.try_make_transport_cps A A;
- v2 <- base.try_make_transport_cps A A;
- Some (fv0 <-- fv;
- Base (v2 (v1 fv0)))%under_lets
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℕ -> (list b) -> (list b)) -> ℕ) -> (list b))%ptype
+ (((ℕ -> (list A) -> (list A)) -> (projT1 args)) -> (list A))%ptype
+ then
+ _ <- ident.unify pattern.ident.List_skipn List_skipn;
+ idc_args0 <- ident.unify pattern.ident.Literal ##(projT2 args);
+ v <- base.try_make_transport_cps A b;
+ v0 <- base.try_make_transport_cps b A;
+ fv <- v0
+ (xs <- reflect_list (v x0);
+ Some
+ (Base
+ (Compilers.reify_list
+ (skipn (let (x1, _) := idc_args0 in x1) xs))));
+ v1 <- base.try_make_transport_cps A A;
+ v2 <- base.try_make_transport_cps A A;
+ Some (fv0 <-- fv;
+ Base (v2 (v1 fv0)))%under_lets
+ else None
| None => None
end
| _ => None
@@ -1359,19 +1516,26 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x1 : option => x1)
with
| Some (_, (_, _), b0, _) =>
- _ <- ident.unify pattern.ident.List_repeat List_repeat;
- v <- base.try_make_transport_cps A b0;
- idc_args0 <- ident.unify pattern.ident.Literal ##(projT2 args);
- v0 <- base.try_make_transport_cps b0 A;
- v1 <- base.try_make_transport_cps A A;
- v2 <- base.try_make_transport_cps A A;
- Some
- (Base
- (v2
- (v1
- (v0
- (Compilers.reify_list
- (repeat (v x) (let (x1, _) := idc_args0 in x1)))))))
+ if
+ type.type_beq base.type base.type.type_beq
+ (((b0 -> ℕ -> (list b0)) -> b0) -> ℕ)%ptype
+ (((A -> ℕ -> (list A)) -> A) -> (projT1 args))%ptype
+ then
+ _ <- ident.unify pattern.ident.List_repeat List_repeat;
+ v <- base.try_make_transport_cps A b0;
+ idc_args0 <- ident.unify pattern.ident.Literal ##(projT2 args);
+ v0 <- base.try_make_transport_cps b0 A;
+ v1 <- base.try_make_transport_cps A A;
+ v2 <- base.try_make_transport_cps A A;
+ Some
+ (Base
+ (v2
+ (v1
+ (v0
+ (Compilers.reify_list
+ (repeat (v x)
+ (let (x1, _) := idc_args0 in x1)))))))
+ else None
| None => None
end
| _ => None
@@ -1389,51 +1553,59 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x1 : option => x1)
with
| Some (_, (_, (_, _)), b0, b) =>
- _ <- ident.unify pattern.ident.List_combine List_combine;
- v <- base.try_make_transport_cps A b0;
- v0 <- base.try_make_transport_cps B b;
- (trA <-- @base.try_make_transport_cps (fun A0 : base.type => option)
- b0 A;
- trB <-- @base.try_make_transport_cps (fun B0 : base.type => option)
- b B;
- return Some (fun v1 : option => trB (trA v1)))%cps option
- (fun tr : option =>
- match tr with
- | Some v1 =>
- x1 <- v1
- (xs <- reflect_list (v x);
- ys <- reflect_list (v0 x0);
- Some
- (Compilers.reify_list
- (map (fun '(x1, y) => (x1, y)%expr_pat)
- (combine xs ys))));
- (trA <-- @base.try_make_transport_cps
- (fun A0 : base.type => expr (list (A0 * B))) A A;
- trB <-- @base.try_make_transport_cps
- (fun B0 : base.type => expr (list (A * B0))) B B;
- return Some (fun v2 : expr (list (A * B)) => trB (trA v2)))%cps
- option
- (fun tr0 : option =>
- match tr0 with
- | Some v2 =>
- (trA <-- @base.try_make_transport_cps
- (fun A0 : base.type => expr (list (A0 * B)))
- A A;
- trB <-- @base.try_make_transport_cps
- (fun B0 : base.type => expr (list (A * B0)))
- B B;
- return Some
- (fun v3 : expr (list (A * B)) =>
- trB (trA v3)))%cps option
- (fun tr' : option =>
- match tr' with
- | Some v3 => Some (Base (v3 (v2 x1)))
- | None => None
- end)
- | None => None
- end)
- | None => None
- end)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((((list b0) -> (list b) -> (list (b0 * b))) -> (list b0)) ->
+ (list b))%ptype
+ ((((list A) -> (list B) -> (list (A * B))) -> (list A)) ->
+ (list B))%ptype
+ then
+ _ <- ident.unify pattern.ident.List_combine List_combine;
+ v <- base.try_make_transport_cps A b0;
+ v0 <- base.try_make_transport_cps B b;
+ (trA <-- @base.try_make_transport_cps
+ (fun A0 : base.type => option) b0 A;
+ trB <-- @base.try_make_transport_cps
+ (fun B0 : base.type => option) b B;
+ return Some (fun v1 : option => trB (trA v1)))%cps option
+ (fun tr : option =>
+ match tr with
+ | Some v1 =>
+ x1 <- v1
+ (xs <- reflect_list (v x);
+ ys <- reflect_list (v0 x0);
+ Some
+ (Compilers.reify_list
+ (map (fun '(x1, y) => (x1, y)%expr_pat)
+ (combine xs ys))));
+ (trA <-- @base.try_make_transport_cps
+ (fun A0 : base.type => expr (list (A0 * B))) A A;
+ trB <-- @base.try_make_transport_cps
+ (fun B0 : base.type => expr (list (A * B0))) B B;
+ return Some (fun v2 : expr (list (A * B)) => trB (trA v2)))%cps
+ option
+ (fun tr0 : option =>
+ match tr0 with
+ | Some v2 =>
+ (trA <-- @base.try_make_transport_cps
+ (fun A0 : base.type =>
+ expr (list (A0 * B))) A A;
+ trB <-- @base.try_make_transport_cps
+ (fun B0 : base.type =>
+ expr (list (A * B0))) B B;
+ return Some
+ (fun v3 : expr (list (A * B)) =>
+ trB (trA v3)))%cps option
+ (fun tr' : option =>
+ match tr' with
+ | Some v3 => Some (Base (v3 (v2 x1)))
+ | None => None
+ end)
+ | None => None
+ end)
+ | None => None
+ end)
+ else None
| None => None
end;;;
Base (#(List_combine)%expr @ x @ x0)%expr_pat)%option
@@ -1448,38 +1620,67 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x1 : option => x1)
with
| Some (_, _, (_, _), (_, b4), b) =>
- _ <- ident.unify pattern.ident.List_map List_map;
- base.try_make_transport_cps A b
- (fun a5 : option =>
- (fa <- (fun (T : Type) (k : option -> T) =>
- match a5 with
- | Some x' =>
- base.try_make_transport_cps B b4
- (fun a6 : option =>
- fa <- (fun (T0 : Type) (k0 : option -> T0) =>
- match a6 with
- | Some x'0 =>
- (return Some
- (fun
- v : expr A ->
- UnderLets (expr B) =>
- x'0 (x' v))) T0 k0
- | None => k0 None
- end);
- k fa)
- | None => k None
- end);
- match fa with
- | Some v =>
- v0 <- base.try_make_transport_cps A b;
- v1 <- base.try_make_transport_cps b4 B;
- (fv <- v1
- (ls <- reflect_list (v0 x0);
- Some
- (Datatypes.list_rect
- (fun
- _ : Datatypes.list
- (expr
+ if
+ type.type_beq base.type base.type.type_beq
+ ((((b -> b4) -> (list b) -> (list b4)) -> b -> b4) -> (list b))%ptype
+ ((((A -> B) -> (list A) -> (list B)) -> A -> B) -> (list A))%ptype
+ then
+ _ <- ident.unify pattern.ident.List_map List_map;
+ base.try_make_transport_cps A b
+ (fun a5 : option =>
+ (fa <- (fun (T : Type) (k : option -> T) =>
+ match a5 with
+ | Some x' =>
+ base.try_make_transport_cps B b4
+ (fun a6 : option =>
+ fa <- (fun (T0 : Type) (k0 : option -> T0) =>
+ match a6 with
+ | Some x'0 =>
+ (return Some
+ (fun
+ v : expr A ->
+ UnderLets (expr B)
+ => x'0 (x' v))) T0 k0
+ | None => k0 None
+ end);
+ k fa)
+ | None => k None
+ end);
+ match fa with
+ | Some v =>
+ v0 <- base.try_make_transport_cps A b;
+ v1 <- base.try_make_transport_cps b4 B;
+ (fv <- v1
+ (ls <- reflect_list (v0 x0);
+ Some
+ (Datatypes.list_rect
+ (fun
+ _ : Datatypes.list
+ (expr
+ (base.subst_default '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 2%positive) b4
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b
+ (PositiveMap.empty
+ base.type))))) =>
+ UnderLets
+ (expr
+ (list
+ (base.subst_default '2
+ (PositiveMap.add
+ (PositiveSet.rev 2%positive)
+ b4
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b
+ (PositiveMap.empty
+ base.type)))))))
+ (Base []%expr_pat)
+ (fun
+ (x1 : expr
(base.subst_default '1
(PositiveMap.add
(PositiveSet.rev
@@ -1488,65 +1689,43 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(PositiveSet.rev
1%positive) b
(PositiveMap.empty
- base.type))))) =>
- UnderLets
- (expr
- (list
- (base.subst_default '2
- (PositiveMap.add
- (PositiveSet.rev 2%positive)
- b4
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type)))))))
- (Base []%expr_pat)
- (fun
- (x1 : expr
- (base.subst_default '1
- (PositiveMap.add
- (PositiveSet.rev 2%positive)
- b4
+ base.type)))))
+ (_ : Datatypes.list
+ (expr
+ (base.subst_default '1
(PositiveMap.add
(PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type)))))
- (_ : Datatypes.list
- (expr
- (base.subst_default '1
- (PositiveMap.add
- (PositiveSet.rev
- 2%positive) b4
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type))))))
- (rec : UnderLets
- (expr
- (list
- (base.subst_default
- '2
- (PositiveMap.add
- (PositiveSet.rev
- 2%positive) b4
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type)))))))
- =>
- (rec' <-- rec;
- fx <-- v x x1;
- Base (fx :: rec')%expr_pat)%under_lets) ls));
- v2 <- base.try_make_transport_cps B B;
- v3 <- base.try_make_transport_cps B B;
- Some (fv0 <-- fv;
- Base (v3 (v2 fv0)))%under_lets)%option
- | None => None
- end)%cps)
+ 2%positive) b4
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b
+ (PositiveMap.empty
+ base.type))))))
+ (rec : UnderLets
+ (expr
+ (list
+ (base.subst_default
+ '2
+ (PositiveMap.add
+ (PositiveSet.rev
+ 2%positive) b4
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b
+ (PositiveMap.empty
+ base.type)))))))
+ =>
+ (rec' <-- rec;
+ fx <-- v x x1;
+ Base (fx :: rec')%expr_pat)%under_lets)
+ ls));
+ v2 <- base.try_make_transport_cps B B;
+ v3 <- base.try_make_transport_cps B B;
+ Some (fv0 <-- fv;
+ Base (v3 (v2 fv0)))%under_lets)%option
+ | None => None
+ end)%cps)
+ else None
| None => None
end;;
None);;;
@@ -1564,37 +1743,44 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x1 : option => x1)
with
| Some (_, (_, _), _, b) =>
- _ <- ident.unify pattern.ident.List_app List_app;
- v <- base.try_make_transport_cps A b;
- v0 <- base.try_make_transport_cps A b;
- v1 <- base.try_make_transport_cps b A;
- fv <- v1
- (ls <- reflect_list (v x);
- Some
- (Datatypes.list_rect
- (fun _ : Datatypes.list (expr b) =>
- UnderLets
- (expr
- (base.subst_default (pattern.base.type.list '1)
- (PositiveMap.add
- (PositiveSet.rev 1%positive) b
- (PositiveMap.empty base.type)))))
- (Base (v0 x0))
- (fun (x1 : expr b) (_ : Datatypes.list (expr b))
- (rec : UnderLets
- (expr
- (base.subst_default
- (pattern.base.type.list '1)
- (PositiveMap.add
- (PositiveSet.rev 1%positive) b
- (PositiveMap.empty base.type)))))
- =>
- (rec' <-- rec;
- Base (x1 :: rec')%expr_pat)%under_lets) ls));
- v2 <- base.try_make_transport_cps A A;
- v3 <- base.try_make_transport_cps A A;
- Some (fv0 <-- fv;
- Base (v3 (v2 fv0)))%under_lets
+ if
+ type.type_beq base.type base.type.type_beq
+ ((((list b) -> (list b) -> (list b)) -> (list b)) -> (list b))%ptype
+ ((((list A) -> (list A) -> (list A)) -> (list A)) -> (list A))%ptype
+ then
+ _ <- ident.unify pattern.ident.List_app List_app;
+ v <- base.try_make_transport_cps A b;
+ v0 <- base.try_make_transport_cps A b;
+ v1 <- base.try_make_transport_cps b A;
+ fv <- v1
+ (ls <- reflect_list (v x);
+ Some
+ (Datatypes.list_rect
+ (fun _ : Datatypes.list (expr b) =>
+ UnderLets
+ (expr
+ (base.subst_default
+ (pattern.base.type.list '1)
+ (PositiveMap.add
+ (PositiveSet.rev 1%positive) b
+ (PositiveMap.empty base.type)))))
+ (Base (v0 x0))
+ (fun (x1 : expr b) (_ : Datatypes.list (expr b))
+ (rec : UnderLets
+ (expr
+ (base.subst_default
+ (pattern.base.type.list '1)
+ (PositiveMap.add
+ (PositiveSet.rev 1%positive) b
+ (PositiveMap.empty base.type)))))
+ =>
+ (rec' <-- rec;
+ Base (x1 :: rec')%expr_pat)%under_lets) ls));
+ v2 <- base.try_make_transport_cps A A;
+ v3 <- base.try_make_transport_cps A A;
+ Some (fv0 <-- fv;
+ Base (v3 (v2 fv0)))%under_lets
+ else None
| None => None
end;;
None);;;
@@ -1609,16 +1795,22 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x0 : option => x0)
with
| Some (_, _, b) =>
- _ <- ident.unify pattern.ident.List_rev List_rev;
- v <- base.try_make_transport_cps A b;
- v0 <- base.try_make_transport_cps b A;
- fv <- v0
- (xs <- reflect_list (v x);
- Some (Base (Compilers.reify_list (rev xs))));
- v1 <- base.try_make_transport_cps A A;
- v2 <- base.try_make_transport_cps A A;
- Some (fv0 <-- fv;
- Base (v2 (v1 fv0)))%under_lets
+ if
+ type.type_beq base.type base.type.type_beq
+ (((list b) -> (list b)) -> (list b))%ptype
+ (((list A) -> (list A)) -> (list A))%ptype
+ then
+ _ <- ident.unify pattern.ident.List_rev List_rev;
+ v <- base.try_make_transport_cps A b;
+ v0 <- base.try_make_transport_cps b A;
+ fv <- v0
+ (xs <- reflect_list (v x);
+ Some (Base (Compilers.reify_list (rev xs))));
+ v1 <- base.try_make_transport_cps A A;
+ v2 <- base.try_make_transport_cps A A;
+ Some (fv0 <-- fv;
+ Base (v2 (v1 fv0)))%under_lets
+ else None
| None => None
end;;
None);;;
@@ -1635,39 +1827,59 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(list A))%ptype option (fun x1 : option => x1)
with
| Some (_, _, (_, _), (_, b4), b) =>
- _ <- ident.unify pattern.ident.List_flat_map List_flat_map;
- base.try_make_transport_cps A b
- (fun a5 : option =>
- (fa <- (fun (T : Type) (k : option -> T) =>
- match a5 with
- | Some x' =>
- base.try_make_transport_cps B b4
- (fun a6 : option =>
- fa <- (fun (T0 : Type) (k0 : option -> T0) =>
- match a6 with
- | Some x'0 =>
- (return Some
- (fun
- v : expr A ->
- UnderLets
- (expr (list B)) =>
- x'0 (x' v))) T0 k0
- | None => k0 None
- end);
- k fa)
- | None => k None
- end);
- match fa with
- | Some v =>
- v0 <- base.try_make_transport_cps A b;
- v1 <- base.try_make_transport_cps b4 B;
- (fv <- v1
- (ls <- reflect_list (v0 x0);
- Some
- (Datatypes.list_rect
- (fun
- _ : Datatypes.list
- (expr
+ if
+ type.type_beq base.type base.type.type_beq
+ ((((b -> (list b4)) -> (list b) -> (list b4)) -> b -> (list b4)) ->
+ (list b))%ptype
+ ((((A -> (list B)) -> (list A) -> (list B)) -> A -> (list B)) ->
+ (list A))%ptype
+ then
+ _ <- ident.unify pattern.ident.List_flat_map List_flat_map;
+ base.try_make_transport_cps A b
+ (fun a5 : option =>
+ (fa <- (fun (T : Type) (k : option -> T) =>
+ match a5 with
+ | Some x' =>
+ base.try_make_transport_cps B b4
+ (fun a6 : option =>
+ fa <- (fun (T0 : Type) (k0 : option -> T0) =>
+ match a6 with
+ | Some x'0 =>
+ (return Some
+ (fun
+ v : expr A ->
+ UnderLets
+ (expr (list B)) =>
+ x'0 (x' v))) T0 k0
+ | None => k0 None
+ end);
+ k fa)
+ | None => k None
+ end);
+ match fa with
+ | Some v =>
+ v0 <- base.try_make_transport_cps A b;
+ v1 <- base.try_make_transport_cps b4 B;
+ (fv <- v1
+ (ls <- reflect_list (v0 x0);
+ Some
+ (Datatypes.list_rect
+ (fun
+ _ : Datatypes.list
+ (expr
+ (base.subst_default '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 2%positive) b4
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b
+ (PositiveMap.empty
+ base.type))))) =>
+ UnderLets (expr (list b4)))
+ (Base []%expr_pat)
+ (fun
+ (x1 : expr
(base.subst_default '1
(PositiveMap.add
(PositiveSet.rev
@@ -1676,43 +1888,31 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(PositiveSet.rev
1%positive) b
(PositiveMap.empty
- base.type))))) =>
- UnderLets (expr (list b4)))
- (Base []%expr_pat)
- (fun
- (x1 : expr
- (base.subst_default '1
- (PositiveMap.add
- (PositiveSet.rev 2%positive)
- b4
+ base.type)))))
+ (_ : Datatypes.list
+ (expr
+ (base.subst_default '1
(PositiveMap.add
(PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type)))))
- (_ : Datatypes.list
- (expr
- (base.subst_default '1
- (PositiveMap.add
- (PositiveSet.rev
- 2%positive) b4
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type))))))
- (rec : UnderLets (expr (list b4))) =>
- (rec' <-- rec;
- fx <-- v x x1;
- Base ($fx ++ rec')%expr)%under_lets) ls));
- v2 <- base.try_make_transport_cps B B;
- v3 <- base.try_make_transport_cps B B;
- Some
- (fv0 <-- fv;
- fv1 <-- do_again (list B) (v2 fv0);
- Base (v3 fv1))%under_lets)%option
- | None => None
- end)%cps)
+ 2%positive) b4
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b
+ (PositiveMap.empty
+ base.type))))))
+ (rec : UnderLets (expr (list b4))) =>
+ (rec' <-- rec;
+ fx <-- v x x1;
+ Base ($fx ++ rec')%expr)%under_lets) ls));
+ v2 <- base.try_make_transport_cps B B;
+ v3 <- base.try_make_transport_cps B B;
+ Some
+ (fv0 <-- fv;
+ fv1 <-- do_again (list B) (v2 fv0);
+ Base (v3 fv1))%under_lets)%option
+ | None => None
+ end)%cps)
+ else None
| None => None
end;;
None);;;
@@ -1731,85 +1931,55 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(list A))%ptype option (fun x1 : option => x1)
with
| Some (_, _, (_, (_, _)), (_, _), b) =>
- _ <- ident.unify pattern.ident.List_partition List_partition;
- base.try_make_transport_cps A b
- (fun a6 : option =>
- (fa <- (fun (T : Type) (k : option -> T) =>
- match a6 with
- | Some x' =>
- (return Some
- (fun v : expr A -> UnderLets (expr bool)
- => x' v)) T (fun fa : option => k fa)
- | None => k None
- end);
- match fa with
- | Some v =>
- v0 <- base.try_make_transport_cps A b;
- base.try_make_transport_cps b A
- (fun a7 : option =>
- fa0 <- (fun (T : Type) (k : option -> T) =>
- match a7 with
- | Some x' =>
- base.try_make_transport_cps b A
- (fun a8 : option =>
- fa0 <- (fun (T0 : Type)
- (k0 : option -> T0) =>
- match a8 with
- | Some x'0 =>
- (return Some
- (fun v1 : option =>
- x'0 (x' v1))) T0
- k0
- | None => k0 None
- end);
- k fa0)
- | None => k None
- end);
- match fa0 with
- | Some v1 =>
- (fv <- v1
- (ls <- reflect_list (v0 x0);
- Some
- (Datatypes.list_rect
- (fun
- _ : Datatypes.list
- (expr
- (base.subst_default
- '1
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type)))) =>
- UnderLets
- (expr
- (list
- (base.subst_default
- '1
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type))) *
- list
- (base.subst_default
- '1
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type))))%etype))
- (Base ([], [])%expr_pat)
- (fun
- (x1 : expr
- (base.subst_default
- '1
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive) b
- (PositiveMap.empty
- base.type))))
- (_ : Datatypes.list
+ if
+ type.type_beq base.type base.type.type_beq
+ ((((b -> bool) -> (list b) -> (list b * list b)%etype) ->
+ b -> bool) -> (list b))%ptype
+ ((((A -> bool) -> (list A) -> (list A * list A)%etype) ->
+ A -> bool) -> (list A))%ptype
+ then
+ _ <- ident.unify pattern.ident.List_partition List_partition;
+ base.try_make_transport_cps A b
+ (fun a6 : option =>
+ (fa <- (fun (T : Type) (k : option -> T) =>
+ match a6 with
+ | Some x' =>
+ (return Some
+ (fun v : expr A -> UnderLets (expr bool)
+ => x' v)) T (fun fa : option => k fa)
+ | None => k None
+ end);
+ match fa with
+ | Some v =>
+ v0 <- base.try_make_transport_cps A b;
+ base.try_make_transport_cps b A
+ (fun a7 : option =>
+ fa0 <- (fun (T : Type) (k : option -> T) =>
+ match a7 with
+ | Some x' =>
+ base.try_make_transport_cps b A
+ (fun a8 : option =>
+ fa0 <- (fun (T0 : Type)
+ (k0 : option -> T0) =>
+ match a8 with
+ | Some x'0 =>
+ (return Some
+ (fun v1 : option
+ => x'0 (x' v1)))
+ T0 k0
+ | None => k0 None
+ end);
+ k fa0)
+ | None => k None
+ end);
+ match fa0 with
+ | Some v1 =>
+ (fv <- v1
+ (ls <- reflect_list (v0 x0);
+ Some
+ (Datatypes.list_rect
+ (fun
+ _ : Datatypes.list
(expr
(base.subst_default
'1
@@ -1817,63 +1987,102 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(PositiveSet.rev
1%positive) b
(PositiveMap.empty
- base.type)))))
- (rec : UnderLets
- (expr
- (list
- (base.subst_default
- '1
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive)
- b
- (PositiveMap.empty
- base.type))) *
- list
- (base.subst_default
- '1
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive)
- b
- (PositiveMap.empty
- base.type))))%etype))
- =>
- (rec' <-- rec;
- fx <-- v x x1;
- Base
- (#(prod_rect)%expr @
- (λ g
- d : expr
- (list
- (base.subst_default
- '1
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive)
- b
- (PositiveMap.empty
- base.type)))),
- (#(bool_rect)%expr @
- (λ _ : expr unit,
- ($x1 :: $g, $d)%expr_pat) @
- (λ _ : expr unit,
- ($g, $x1 :: $d)%expr_pat) @
- $fx)%expr_pat)%expr @ rec')%expr_pat)%under_lets)
- ls));
- base.try_make_transport_cps A A
- (fun a8 : option =>
- (fa1 <- (fun (T : Type) (k : option -> T) =>
- match a8 with
- | Some x' =>
- base.try_make_transport_cps A A
- (fun a9 : option =>
- fa1 <- (fun (T0 : Type)
- (k0 : option -> T0) =>
- match a9 with
- | Some x'0 =>
- (return Some
- (fun
+ base.type))))
+ =>
+ UnderLets
+ (expr
+ (list
+ (base.subst_default
+ '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b
+ (PositiveMap.empty
+ base.type))) *
+ list
+ (base.subst_default
+ '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b
+ (PositiveMap.empty
+ base.type))))%etype))
+ (Base ([], [])%expr_pat)
+ (fun
+ (x1 : expr
+ (base.subst_default
+ '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b
+ (PositiveMap.empty
+ base.type))))
+ (_ : Datatypes.list
+ (expr
+ (base.subst_default
+ '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b
+ (PositiveMap.empty
+ base.type)))))
+ (rec : UnderLets
+ (expr
+ (list
+ (base.subst_default
+ '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive)
+ b
+ (PositiveMap.empty
+ base.type))) *
+ list
+ (base.subst_default
+ '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive)
+ b
+ (PositiveMap.empty
+ base.type))))%etype))
+ =>
+ (rec' <-- rec;
+ fx <-- v x x1;
+ Base
+ (#(prod_rect)%expr @
+ (λ g
+ d : expr
+ (list
+ (base.subst_default
+ '1
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive)
+ b
+ (PositiveMap.empty
+ base.type)))),
+ (#(bool_rect)%expr @
+ (λ _ : expr unit,
+ ($x1 :: $g, $d)%expr_pat) @
+ (λ _ : expr unit,
+ ($g, $x1 :: $d)%expr_pat) @
+ $fx)%expr_pat)%expr @ rec')%expr_pat)%under_lets)
+ ls));
+ base.try_make_transport_cps A A
+ (fun a8 : option =>
+ (fa1 <- (fun (T : Type) (k : option -> T) =>
+ match a8 with
+ | Some x' =>
+ base.try_make_transport_cps A A
+ (fun a9 : option =>
+ fa1 <- (fun (T0 : Type)
+ (k0 : option -> T0)
+ =>
+ match a9 with
+ | Some x'0 =>
+ (return Some
+ (fun
v2 :
expr
(list A *
@@ -1881,31 +2090,32 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
=>
x'0
(x' v2)))
- T0 k0
- | None => k0 None
- end);
- k fa1)
- | None => k None
- end);
- match fa1 with
- | Some v2 =>
- base.try_make_transport_cps A A
- (fun a9 : option =>
- fa2 <- (fun (T : Type) (k : option -> T)
- =>
- match a9 with
- | Some x' =>
- base.try_make_transport_cps
- A A
- (fun a10 : option =>
- fa2 <- (fun (T0 : Type)
- (k0 : option ->
- T0) =>
- match a10 with
- | Some x'0 =>
- (return
- Some
- (fun
+ T0 k0
+ | None => k0 None
+ end);
+ k fa1)
+ | None => k None
+ end);
+ match fa1 with
+ | Some v2 =>
+ base.try_make_transport_cps A A
+ (fun a9 : option =>
+ fa2 <- (fun (T : Type)
+ (k : option -> T) =>
+ match a9 with
+ | Some x' =>
+ base.try_make_transport_cps
+ A A
+ (fun a10 : option =>
+ fa2 <- (fun (T0 : Type)
+ (k0 :
+ option -> T0)
+ =>
+ match a10 with
+ | Some x'0 =>
+ (return
+ Some
+ (fun
v3 :
expr
(list A *
@@ -1913,28 +2123,30 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
=>
x'0
(x' v3)))
- T0 k0
- | None => k0 None
- end);
- k fa2)
- | None => k None
- end);
- match fa2 with
- | Some v3 =>
- Some
- (fv0 <-- fv;
- fv1 <-- do_again
- (list A * list A)
- (v2 fv0);
- Base (v3 fv1))%under_lets
- | None => None
- end)
- | None => None
- end)%cps))%option
- | None => None
- end)
- | None => None
- end)%cps)
+ T0 k0
+ | None =>
+ k0 None
+ end);
+ k fa2)
+ | None => k None
+ end);
+ match fa2 with
+ | Some v3 =>
+ Some
+ (fv0 <-- fv;
+ fv1 <-- do_again
+ (list A * list A)
+ (v2 fv0);
+ Base (v3 fv1))%under_lets
+ | None => None
+ end)
+ | None => None
+ end)%cps))%option
+ | None => None
+ end)
+ | None => None
+ end)%cps)
+ else None
| None => None
end;;
None);;;
@@ -1954,64 +2166,92 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(list B))%ptype option (fun x2 : option => x2)
with
| Some (_, (_, _), (_, (_, _)), (_, (_, _)), b0, b) =>
- _ <- ident.unify pattern.ident.List_fold_right List_fold_right;
- base.try_make_transport_cps B b
- (fun a9 : option =>
- (fa <- (fun (T : Type) (k : option -> T) =>
- match a9 with
- | Some x' =>
- base.try_make_transport_cps A b0
- (fun a10 : option =>
- fa <- (fun (T0 : Type) (k0 : option -> T0) =>
- match a10 with
- | Some x'0 =>
- base.try_make_transport_cps A b0
- (fun a11 : option =>
- fa <- (fun (T1 : Type)
- (k1 : option -> T1) =>
- match a11 with
- | Some x'1 =>
- (return Some
- (fun
- v :
- expr b ->
- expr A ->
- UnderLets
- (expr A)
- =>
- x'1 (x'0 v)))
- T1 k1
- | None => k1 None
- end);
- k0 fa)
- | None => k0 None
- end);
- fa0 <- (fun (T0 : Type) (k0 : option -> T0) =>
- match fa with
+ if
+ type.type_beq base.type base.type.type_beq
+ (((((b -> b0 -> b0) -> b0 -> (list b) -> b0) -> b -> b0 -> b0) ->
+ b0) -> (list b))%ptype
+ (((((B -> A -> A) -> A -> (list B) -> A) -> B -> A -> A) -> A) ->
+ (list B))%ptype
+ then
+ _ <- ident.unify pattern.ident.List_fold_right List_fold_right;
+ base.try_make_transport_cps B b
+ (fun a9 : option =>
+ (fa <- (fun (T : Type) (k : option -> T) =>
+ match a9 with
+ | Some x' =>
+ base.try_make_transport_cps A b0
+ (fun a10 : option =>
+ fa <- (fun (T0 : Type) (k0 : option -> T0) =>
+ match a10 with
| Some x'0 =>
- (return Some
- (fun
- v : expr B ->
- expr A ->
- UnderLets (expr A)
- => x'0 (x' v))) T0 k0
+ base.try_make_transport_cps A b0
+ (fun a11 : option =>
+ fa <- (fun (T1 : Type)
+ (k1 : option -> T1) =>
+ match a11 with
+ | Some x'1 =>
+ (return Some
+ (fun
+ v :
+ expr b ->
+ expr A ->
+ UnderLets
+ (expr A)
+ =>
+ x'1 (x'0 v)))
+ T1 k1
+ | None => k1 None
+ end);
+ k0 fa)
| None => k0 None
end);
- k fa0)
- | None => k None
- end);
- match fa with
- | Some v =>
- v0 <- base.try_make_transport_cps A b0;
- v1 <- base.try_make_transport_cps B b;
- v2 <- base.try_make_transport_cps b0 A;
- (fv <- v2
- (ls <- reflect_list (v1 x1);
- Some
- (Datatypes.list_rect
- (fun
- _ : Datatypes.list
- (expr
+ fa0 <- (fun (T0 : Type) (k0 : option -> T0) =>
+ match fa with
+ | Some x'0 =>
+ (return Some
+ (fun
+ v : expr B ->
+ expr A ->
+ UnderLets (expr A)
+ => x'0 (x' v))) T0 k0
+ | None => k0 None
+ end);
+ k fa0)
+ | None => k None
+ end);
+ match fa with
+ | Some v =>
+ v0 <- base.try_make_transport_cps A b0;
+ v1 <- base.try_make_transport_cps B b;
+ v2 <- base.try_make_transport_cps b0 A;
+ (fv <- v2
+ (ls <- reflect_list (v1 x1);
+ Some
+ (Datatypes.list_rect
+ (fun
+ _ : Datatypes.list
+ (expr
+ (base.subst_default '2
+ (PositiveMap.add
+ (PositiveSet.rev
+ 2%positive) b
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b0
+ (PositiveMap.empty
+ base.type))))) =>
+ UnderLets
+ (expr
+ (base.subst_default '1
+ (PositiveMap.add
+ (PositiveSet.rev 2%positive) b
+ (PositiveMap.add
+ (PositiveSet.rev 1%positive)
+ b0
+ (PositiveMap.empty base.type))))))
+ (Base (v0 x0))
+ (fun
+ (x2 : expr
(base.subst_default '2
(PositiveMap.add
(PositiveSet.rev
@@ -2020,58 +2260,39 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(PositiveSet.rev
1%positive) b0
(PositiveMap.empty
- base.type))))) =>
- UnderLets
- (expr
- (base.subst_default '1
- (PositiveMap.add
- (PositiveSet.rev 2%positive) b
- (PositiveMap.add
- (PositiveSet.rev 1%positive)
- b0
- (PositiveMap.empty base.type))))))
- (Base (v0 x0))
- (fun
- (x2 : expr
- (base.subst_default '2
- (PositiveMap.add
- (PositiveSet.rev 2%positive)
- b
+ base.type)))))
+ (_ : Datatypes.list
+ (expr
+ (base.subst_default '2
(PositiveMap.add
(PositiveSet.rev
- 1%positive) b0
- (PositiveMap.empty
- base.type)))))
- (_ : Datatypes.list
- (expr
- (base.subst_default '2
- (PositiveMap.add
- (PositiveSet.rev
- 2%positive) b
+ 2%positive) b
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b0
+ (PositiveMap.empty
+ base.type))))))
+ (rec : UnderLets
+ (expr
+ (base.subst_default
+ '1
(PositiveMap.add
(PositiveSet.rev
- 1%positive) b0
- (PositiveMap.empty
- base.type))))))
- (rec : UnderLets
- (expr
- (base.subst_default '1
- (PositiveMap.add
- (PositiveSet.rev
- 2%positive) b
- (PositiveMap.add
- (PositiveSet.rev
- 1%positive) b0
- (PositiveMap.empty
- base.type))))))
- => (rec' <-- rec;
- v x x2 rec')%under_lets) ls));
- v3 <- base.try_make_transport_cps A A;
- v4 <- base.try_make_transport_cps A A;
- Some (fv0 <-- fv;
- Base (v4 (v3 fv0)))%under_lets)%option
- | None => None
- end)%cps)
+ 2%positive) b
+ (PositiveMap.add
+ (PositiveSet.rev
+ 1%positive) b0
+ (PositiveMap.empty
+ base.type))))))
+ => (rec' <-- rec;
+ v x x2 rec')%under_lets) ls));
+ v3 <- base.try_make_transport_cps A A;
+ v4 <- base.try_make_transport_cps A A;
+ Some (fv0 <-- fv;
+ Base (v4 (v3 fv0)))%under_lets)%option
+ | None => None
+ end)%cps)
+ else None
| None => None
end;;
None);;;
@@ -2095,58 +2316,67 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
T -> T) -> (list T))%ptype option (fun x2 : option => x2)
with
| Some (_, (_, _, (_, _)), _, (_, _), b) =>
- _ <- ident.unify pattern.ident.List_update_nth List_update_nth;
- idc_args0 <- ident.unify pattern.ident.Literal ##(projT2 args);
- base.try_make_transport_cps T b
- (fun a7 : option =>
- (fa <- (fun (T0 : Type) (k : option -> T0) =>
- match a7 with
- | Some x' =>
- base.try_make_transport_cps T b
- (fun a8 : option =>
- fa <- (fun (T1 : Type) (k0 : option -> T1) =>
- match a8 with
- | Some x'0 =>
- (return Some
- (fun
- v : expr T ->
- UnderLets (expr T)
- => x'0 (x' v))) T1 k0
- | None => k0 None
- end);
- k fa)
- | None => k None
- end);
- match fa with
- | Some v =>
- v0 <- base.try_make_transport_cps T b;
- v1 <- base.try_make_transport_cps b T;
- (fv <- v1
- (ls <- reflect_list (v0 x1);
- Some
- (retv <---- update_nth
- (let (x2, _) := idc_args0 in
- x2)
- (fun
- x2 : UnderLets
- (expr
- (base.subst_default
- '1
- (PositiveMap.add
- (PositiveSet.rev
+ if
+ type.type_beq base.type base.type.type_beq
+ ((((ℕ -> (b -> b) -> (list b) -> (list b)) -> ℕ) -> b -> b) ->
+ (list b))%ptype
+ ((((ℕ -> (T -> T) -> (list T) -> (list T)) -> (projT1 args)) ->
+ T -> T) -> (list T))%ptype
+ then
+ _ <- ident.unify pattern.ident.List_update_nth List_update_nth;
+ idc_args0 <- ident.unify pattern.ident.Literal ##(projT2 args);
+ base.try_make_transport_cps T b
+ (fun a7 : option =>
+ (fa <- (fun (T0 : Type) (k : option -> T0) =>
+ match a7 with
+ | Some x' =>
+ base.try_make_transport_cps T b
+ (fun a8 : option =>
+ fa <- (fun (T1 : Type) (k0 : option -> T1) =>
+ match a8 with
+ | Some x'0 =>
+ (return Some
+ (fun
+ v : expr T ->
+ UnderLets
+ (expr T) =>
+ x'0 (x' v))) T1 k0
+ | None => k0 None
+ end);
+ k fa)
+ | None => k None
+ end);
+ match fa with
+ | Some v =>
+ v0 <- base.try_make_transport_cps T b;
+ v1 <- base.try_make_transport_cps b T;
+ (fv <- v1
+ (ls <- reflect_list (v0 x1);
+ Some
+ (retv <---- update_nth
+ (let (x2, _) := idc_args0 in
+ x2)
+ (fun
+ x2 : UnderLets
+ (expr
+ (base.subst_default
+ '1
+ (PositiveMap.add
+ (PositiveSet.rev
1%positive)
- b
- (PositiveMap.empty
+ b
+ (PositiveMap.empty
base.type))))
- => x3 <-- x2;
- v x0 x3) (map Base ls);
- Base (Compilers.reify_list retv))%under_lets);
- v2 <- base.try_make_transport_cps T T;
- v3 <- base.try_make_transport_cps T T;
- Some (fv0 <-- fv;
- Base (v3 (v2 fv0)))%under_lets)%option
- | None => None
- end)%cps)
+ => x3 <-- x2;
+ v x0 x3) (map Base ls);
+ Base (Compilers.reify_list retv))%under_lets);
+ v2 <- base.try_make_transport_cps T T;
+ v3 <- base.try_make_transport_cps T T;
+ Some (fv0 <-- fv;
+ Base (v3 (v2 fv0)))%under_lets)%option
+ | None => None
+ end)%cps)
+ else None
| None => None
end
| _ => None
@@ -2167,20 +2397,27 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(projT1 args))%ptype option (fun x2 : option => x2)
with
| Some (_, (_, (_, _)), _, b0, _) =>
- _ <- ident.unify pattern.ident.List_nth_default
- List_nth_default;
- v <- base.try_make_transport_cps T b0;
- v0 <- base.try_make_transport_cps T b0;
- idc_args0 <- ident.unify pattern.ident.Literal ##(projT2 args);
- v1 <- base.try_make_transport_cps b0 T;
- x2 <- v1
- (ls <- reflect_list (v0 x0);
- Some
- (nth_default (v x) ls
- (let (x2, _) := idc_args0 in x2)));
- v2 <- base.try_make_transport_cps T T;
- v3 <- base.try_make_transport_cps T T;
- Some (Base (v3 (v2 x2)))
+ if
+ type.type_beq base.type base.type.type_beq
+ ((((b0 -> (list b0) -> ℕ -> b0) -> b0) -> (list b0)) -> ℕ)%ptype
+ ((((T -> (list T) -> ℕ -> T) -> T) -> (list T)) ->
+ (projT1 args))%ptype
+ then
+ _ <- ident.unify pattern.ident.List_nth_default
+ List_nth_default;
+ v <- base.try_make_transport_cps T b0;
+ v0 <- base.try_make_transport_cps T b0;
+ idc_args0 <- ident.unify pattern.ident.Literal ##(projT2 args);
+ v1 <- base.try_make_transport_cps b0 T;
+ x2 <- v1
+ (ls <- reflect_list (v0 x0);
+ Some
+ (nth_default (v x) ls
+ (let (x2, _) := idc_args0 in x2)));
+ v2 <- base.try_make_transport_cps T T;
+ v3 <- base.try_make_transport_cps T T;
+ Some (Base (v3 (v2 x2)))
+ else None
| None => None
end
| _ => None
@@ -2201,14 +2438,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##((let (x1, _) := idc_args in x1) +
- (let (x1, _) := idc_args0 in x1))%Z)%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##((let (x1, _) := idc_args in x1) +
+ (let (x1, _) := idc_args0 in x1))%Z)%expr)
+ else None
| None => None
end
| _ => None
@@ -2231,14 +2473,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##((let (x1, _) := idc_args in x1) *
- (let (x1, _) := idc_args0 in x1))%Z)%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##((let (x1, _) := idc_args in x1) *
+ (let (x1, _) := idc_args0 in x1))%Z)%expr)
+ else None
| None => None
end
| _ => None
@@ -2261,14 +2508,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##((let (x1, _) := idc_args in x1)
- ^ (let (x1, _) := idc_args0 in x1)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##((let (x1, _) := idc_args in x1)
+ ^ (let (x1, _) := idc_args0 in x1)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2291,14 +2543,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##((let (x1, _) := idc_args in x1) -
- (let (x1, _) := idc_args0 in x1))%Z)%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##((let (x1, _) := idc_args in x1) -
+ (let (x1, _) := idc_args0 in x1))%Z)%expr)
+ else None
| None => None
end
| _ => None
@@ -2317,8 +2574,11 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x0 : option => x0)
with
| Some _ =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- Some (Base (##(- (let (x0, _) := idc_args in x0))%Z)%expr)
+ if type.type_beq base.type base.type.type_beq ℤ (projT1 args)
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ Some (Base (##(- (let (x0, _) := idc_args in x0))%Z)%expr)
+ else None
| None => None
end
| _ => None
@@ -2339,14 +2599,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##((let (x1, _) := idc_args in x1) /
- (let (x1, _) := idc_args0 in x1))%Z)%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##((let (x1, _) := idc_args in x1) /
+ (let (x1, _) := idc_args0 in x1))%Z)%expr)
+ else None
| None => None
end
| _ => None
@@ -2369,14 +2634,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##((let (x1, _) := idc_args in x1)
- mod (let (x1, _) := idc_args0 in x1))%Z)%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##((let (x1, _) := idc_args in x1)
+ mod (let (x1, _) := idc_args0 in x1))%Z)%expr)
+ else None
| None => None
end
| _ => None
@@ -2395,8 +2665,11 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x0 : option => x0)
with
| Some _ =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- Some (Base (##(Z.log2 (let (x0, _) := idc_args in x0)))%expr)
+ if type.type_beq base.type base.type.type_beq ℤ (projT1 args)
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ Some (Base (##(Z.log2 (let (x0, _) := idc_args in x0)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2413,9 +2686,12 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x0 : option => x0)
with
| Some _ =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- Some
- (Base (##(Z.log2_up (let (x0, _) := idc_args in x0)))%expr)
+ if type.type_beq base.type base.type.type_beq ℤ (projT1 args)
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ Some
+ (Base (##(Z.log2_up (let (x0, _) := idc_args in x0)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2436,14 +2712,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##((let (x1, _) := idc_args in x1) =?
- (let (x1, _) := idc_args0 in x1)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##((let (x1, _) := idc_args in x1) =?
+ (let (x1, _) := idc_args0 in x1)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2466,14 +2747,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##((let (x1, _) := idc_args in x1) <=?
- (let (x1, _) := idc_args0 in x1)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##((let (x1, _) := idc_args in x1) <=?
+ (let (x1, _) := idc_args0 in x1)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2496,14 +2782,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##((let (x1, _) := idc_args in x1) >=?
- (let (x1, _) := idc_args0 in x1)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##((let (x1, _) := idc_args in x1) >=?
+ (let (x1, _) := idc_args0 in x1)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2522,8 +2813,12 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x0 : option => x0)
with
| Some _ =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- Some (Base (##(Z.of_nat (let (x0, _) := idc_args in x0)))%expr)
+ if type.type_beq base.type base.type.type_beq ℕ (projT1 args)
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ Some
+ (Base (##(Z.of_nat (let (x0, _) := idc_args in x0)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2540,8 +2835,12 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x0 : option => x0)
with
| Some _ =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- Some (Base (##(Z.to_nat (let (x0, _) := idc_args in x0)))%expr)
+ if type.type_beq base.type base.type.type_beq ℤ (projT1 args)
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ Some
+ (Base (##(Z.to_nat (let (x0, _) := idc_args in x0)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2562,14 +2861,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##(Z.shiftr (let (x1, _) := idc_args in x1)
- (let (x1, _) := idc_args0 in x1)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##(Z.shiftr (let (x1, _) := idc_args in x1)
+ (let (x1, _) := idc_args0 in x1)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2592,14 +2896,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##(Z.shiftl (let (x1, _) := idc_args in x1)
- (let (x1, _) := idc_args0 in x1)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##(Z.shiftl (let (x1, _) := idc_args in x1)
+ (let (x1, _) := idc_args0 in x1)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2622,14 +2931,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##(Z.land (let (x1, _) := idc_args in x1)
- (let (x1, _) := idc_args0 in x1)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##(Z.land (let (x1, _) := idc_args in x1)
+ (let (x1, _) := idc_args0 in x1)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2652,14 +2966,19 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##(Z.lor (let (x1, _) := idc_args in x1)
- (let (x1, _) := idc_args0 in x1)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##(Z.lor (let (x1, _) := idc_args in x1)
+ (let (x1, _) := idc_args0 in x1)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2678,10 +2997,13 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x0 : option => x0)
with
| Some _ =>
- idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
- Some
- (Base
- (##(Definitions.Z.bneg (let (x0, _) := idc_args in x0)))%expr)
+ if type.type_beq base.type base.type.type_beq ℤ (projT1 args)
+ then
+ idc_args <- ident.unify pattern.ident.Literal ##(projT2 args);
+ Some
+ (Base
+ (##(Definitions.Z.bneg (let (x0, _) := idc_args in x0)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2702,15 +3024,20 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##(Definitions.Z.lnot_modulo
- (let (x1, _) := idc_args in x1)
- (let (x1, _) := idc_args0 in x1)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##(Definitions.Z.lnot_modulo
+ (let (x1, _) := idc_args in x1)
+ (let (x1, _) := idc_args0 in x1)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2736,21 +3063,27 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x2 : option => x2)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (let
- '(a1, b1) :=
- Definitions.Z.mul_split
- (let (x2, _) := idc_args in x2)
- (let (x2, _) := idc_args0 in x2)
- (let (x2, _) := idc_args1 in x2) in
- ((##a1)%expr, (##b1)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> (projT1 args0)) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (let
+ '(a1, b1) :=
+ Definitions.Z.mul_split
+ (let (x2, _) := idc_args in x2)
+ (let (x2, _) := idc_args0 in x2)
+ (let (x2, _) := idc_args1 in x2) in
+ ((##a1)%expr, (##b1)%expr)%expr_pat))
+ else None
| None => None
end
| _ => None
@@ -2778,21 +3111,27 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x2 : option => x2)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (let
- '(a1, b1) :=
- Definitions.Z.add_get_carry_full
- (let (x2, _) := idc_args in x2)
- (let (x2, _) := idc_args0 in x2)
- (let (x2, _) := idc_args1 in x2) in
- ((##a1)%expr, (##b1)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> (projT1 args0)) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (let
+ '(a1, b1) :=
+ Definitions.Z.add_get_carry_full
+ (let (x2, _) := idc_args in x2)
+ (let (x2, _) := idc_args0 in x2)
+ (let (x2, _) := idc_args1 in x2) in
+ ((##a1)%expr, (##b1)%expr)%expr_pat))
+ else None
| None => None
end
| _ => None
@@ -2820,18 +3159,24 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x2 : option => x2)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##(Definitions.Z.add_with_carry
- (let (x2, _) := idc_args in x2)
- (let (x2, _) := idc_args0 in x2)
- (let (x2, _) := idc_args1 in x2)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> (projT1 args0)) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##(Definitions.Z.add_with_carry
+ (let (x2, _) := idc_args in x2)
+ (let (x2, _) := idc_args0 in x2)
+ (let (x2, _) := idc_args1 in x2)))%expr)
+ else None
| None => None
end
| _ => None
@@ -2864,24 +3209,31 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args1 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args2 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (let
- '(a2, b2) :=
- Definitions.Z.add_with_get_carry_full
- (let (x3, _) := idc_args in x3)
- (let (x3, _) := idc_args0 in x3)
- (let (x3, _) := idc_args1 in x3)
- (let (x3, _) := idc_args2 in x3) in
- ((##a2)%expr, (##b2)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args2) -> (projT1 args1)) ->
+ (projT1 args0)) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args2);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args2 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (let
+ '(a2, b2) :=
+ Definitions.Z.add_with_get_carry_full
+ (let (x3, _) := idc_args in x3)
+ (let (x3, _) := idc_args0 in x3)
+ (let (x3, _) := idc_args1 in x3)
+ (let (x3, _) := idc_args2 in x3) in
+ ((##a2)%expr, (##b2)%expr)%expr_pat))
+ else None
| None => None
end
| _ => None
@@ -2911,21 +3263,27 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x2 : option => x2)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (let
- '(a1, b1) :=
- Definitions.Z.sub_get_borrow_full
- (let (x2, _) := idc_args in x2)
- (let (x2, _) := idc_args0 in x2)
- (let (x2, _) := idc_args1 in x2) in
- ((##a1)%expr, (##b1)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> (projT1 args0)) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (let
+ '(a1, b1) :=
+ Definitions.Z.sub_get_borrow_full
+ (let (x2, _) := idc_args in x2)
+ (let (x2, _) := idc_args0 in x2)
+ (let (x2, _) := idc_args1 in x2) in
+ ((##a1)%expr, (##b1)%expr)%expr_pat))
+ else None
| None => None
end
| _ => None
@@ -2958,24 +3316,31 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args1 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args2 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (let
- '(a2, b2) :=
- Definitions.Z.sub_with_get_borrow_full
- (let (x3, _) := idc_args in x3)
- (let (x3, _) := idc_args0 in x3)
- (let (x3, _) := idc_args1 in x3)
- (let (x3, _) := idc_args2 in x3) in
- ((##a2)%expr, (##b2)%expr)%expr_pat))
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args2) -> (projT1 args1)) ->
+ (projT1 args0)) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args2);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args2 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (let
+ '(a2, b2) :=
+ Definitions.Z.sub_with_get_borrow_full
+ (let (x3, _) := idc_args in x3)
+ (let (x3, _) := idc_args0 in x3)
+ (let (x3, _) := idc_args1 in x3)
+ (let (x3, _) := idc_args2 in x3) in
+ ((##a2)%expr, (##b2)%expr)%expr_pat))
+ else None
| None => None
end
| _ => None
@@ -3005,18 +3370,24 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x2 : option => x2)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##(Definitions.Z.zselect
- (let (x2, _) := idc_args in x2)
- (let (x2, _) := idc_args0 in x2)
- (let (x2, _) := idc_args1 in x2)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> (projT1 args0)) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##(Definitions.Z.zselect
+ (let (x2, _) := idc_args in x2)
+ (let (x2, _) := idc_args0 in x2)
+ (let (x2, _) := idc_args1 in x2)))%expr)
+ else None
| None => None
end
| _ => None
@@ -3044,18 +3415,24 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x2 : option => x2)
with
| Some (_, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##(Definitions.Z.add_modulo
- (let (x2, _) := idc_args in x2)
- (let (x2, _) := idc_args0 in x2)
- (let (x2, _) := idc_args1 in x2)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> (projT1 args0)) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##(Definitions.Z.add_modulo
+ (let (x2, _) := idc_args in x2)
+ (let (x2, _) := idc_args0 in x2)
+ (let (x2, _) := idc_args1 in x2)))%expr)
+ else None
| None => None
end
| _ => None
@@ -3088,21 +3465,28 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x3 : option => x3)
with
| Some (_, _, _, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- idc_args1 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args2 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##(Definitions.Z.rshi
- (let (x3, _) := idc_args in x3)
- (let (x3, _) := idc_args0 in x3)
- (let (x3, _) := idc_args1 in x3)
- (let (x3, _) := idc_args2 in x3)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args2) -> (projT1 args1)) ->
+ (projT1 args0)) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args2);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args2 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##(Definitions.Z.rshi
+ (let (x3, _) := idc_args in x3)
+ (let (x3, _) := idc_args0 in x3)
+ (let (x3, _) := idc_args1 in x3)
+ (let (x3, _) := idc_args2 in x3)))%expr)
+ else None
| None => None
end
| _ => None
@@ -3129,14 +3513,20 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(fun x1 : option => x1)
with
| Some (_, _) =>
- idc_args <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- idc_args0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (##(Definitions.Z.cc_m (let (x1, _) := idc_args in x1)
- (let (x1, _) := idc_args0 in x1)))%expr)
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args0) -> (projT1 args))%ptype
+ then
+ idc_args <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ Some
+ (Base
+ (##(Definitions.Z.cc_m
+ (let (x1, _) := idc_args in x1)
+ (let (x1, _) := idc_args0 in x1)))%expr)
+ else None
| None => None
end
| _ => None
@@ -3200,17 +3590,34 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x4 : option => x4)
with
| Some (_, _, (_, (_, _, _)), (_, (_, (_, _)), _, _), _) =>
- _ <- ident.unify pattern.ident.pair pair;
- _ <- ident.unify pattern.ident.pair pair;
- idc_args1 <- ident.unify pattern.ident.Literal ##(projT2 args1);
- idc_args2 <- ident.unify pattern.ident.Literal ##(projT2 args0);
- idc_args3 <- ident.unify pattern.ident.Literal ##(projT2 args);
- Some
- (Base
- (##(fancy.interp (invert_Some (to_fancy fancy_selc))
- (let (x4, _) := idc_args1 in x4,
- let (x4, _) := idc_args2 in x4,
- let (x4, _) := idc_args3 in x4)%core))%expr)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((((ℤ * ℤ)%etype -> ℤ -> (ℤ * ℤ * ℤ)%etype) ->
+ ((ℤ -> ℤ -> (ℤ * ℤ)%etype) -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((let (x4, _) := args3 in x4) ->
+ (let (_, y) := args3 in y) ->
+ ((let (x4, _) := args3 in x4) *
+ (let (_, y) := args3 in y))%etype) ->
+ (((let (x4, _) := args2 in x4) ->
+ (let (_, y) := args2 in y) ->
+ ((let (x4, _) := args2 in x4) *
+ (let (_, y) := args2 in y))%etype) -> (projT1 args1)) ->
+ (projT1 args0)) -> (projT1 args))%ptype
+ then
+ _ <- ident.unify pattern.ident.pair pair;
+ _ <- ident.unify pattern.ident.pair pair;
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args2 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args3 <- ident.unify pattern.ident.Literal ##(projT2 args);
+ Some
+ (Base
+ (##(fancy.interp (invert_Some (to_fancy fancy_selc))
+ (let (x4, _) := idc_args1 in x4,
+ let (x4, _) := idc_args2 in x4,
+ let (x4, _) := idc_args3 in x4)%core))%expr)
+ else None
| None => None
end
| (@expr.Ident _ _ _ t idc @
@@ -3291,17 +3698,34 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x4 : option => x4)
with
| Some (_, _, (_, (_, _, _)), (_, (_, (_, _)), _, _), _) =>
- _ <- ident.unify pattern.ident.pair pair;
- _ <- ident.unify pattern.ident.pair pair;
- idc_args1 <- ident.unify pattern.ident.Literal ##(projT2 args1);
- idc_args2 <- ident.unify pattern.ident.Literal ##(projT2 args0);
- idc_args3 <- ident.unify pattern.ident.Literal ##(projT2 args);
- Some
- (Base
- (##(fancy.interp (invert_Some (to_fancy fancy_sell))
- (let (x4, _) := idc_args1 in x4,
- let (x4, _) := idc_args2 in x4,
- let (x4, _) := idc_args3 in x4)%core))%expr)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((((ℤ * ℤ)%etype -> ℤ -> (ℤ * ℤ * ℤ)%etype) ->
+ ((ℤ -> ℤ -> (ℤ * ℤ)%etype) -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((let (x4, _) := args3 in x4) ->
+ (let (_, y) := args3 in y) ->
+ ((let (x4, _) := args3 in x4) *
+ (let (_, y) := args3 in y))%etype) ->
+ (((let (x4, _) := args2 in x4) ->
+ (let (_, y) := args2 in y) ->
+ ((let (x4, _) := args2 in x4) *
+ (let (_, y) := args2 in y))%etype) -> (projT1 args1)) ->
+ (projT1 args0)) -> (projT1 args))%ptype
+ then
+ _ <- ident.unify pattern.ident.pair pair;
+ _ <- ident.unify pattern.ident.pair pair;
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args2 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args3 <- ident.unify pattern.ident.Literal ##(projT2 args);
+ Some
+ (Base
+ (##(fancy.interp (invert_Some (to_fancy fancy_sell))
+ (let (x4, _) := idc_args1 in x4,
+ let (x4, _) := idc_args2 in x4,
+ let (x4, _) := idc_args3 in x4)%core))%expr)
+ else None
| None => None
end
| (@expr.Ident _ _ _ t idc @
@@ -3379,17 +3803,34 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
option (fun x4 : option => x4)
with
| Some (_, _, (_, (_, _, _)), (_, (_, (_, _)), _, _), _) =>
- _ <- ident.unify pattern.ident.pair pair;
- _ <- ident.unify pattern.ident.pair pair;
- idc_args1 <- ident.unify pattern.ident.Literal ##(projT2 args1);
- idc_args2 <- ident.unify pattern.ident.Literal ##(projT2 args0);
- idc_args3 <- ident.unify pattern.ident.Literal ##(projT2 args);
- Some
- (Base
- (##(fancy.interp (invert_Some (to_fancy fancy_addm))
- (let (x4, _) := idc_args1 in x4,
- let (x4, _) := idc_args2 in x4,
- let (x4, _) := idc_args3 in x4)%core))%expr)
+ if
+ type.type_beq base.type base.type.type_beq
+ ((((ℤ * ℤ)%etype -> ℤ -> (ℤ * ℤ * ℤ)%etype) ->
+ ((ℤ -> ℤ -> (ℤ * ℤ)%etype) -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((let (x4, _) := args3 in x4) ->
+ (let (_, y) := args3 in y) ->
+ ((let (x4, _) := args3 in x4) *
+ (let (_, y) := args3 in y))%etype) ->
+ (((let (x4, _) := args2 in x4) ->
+ (let (_, y) := args2 in y) ->
+ ((let (x4, _) := args2 in x4) *
+ (let (_, y) := args2 in y))%etype) -> (projT1 args1)) ->
+ (projT1 args0)) -> (projT1 args))%ptype
+ then
+ _ <- ident.unify pattern.ident.pair pair;
+ _ <- ident.unify pattern.ident.pair pair;
+ idc_args1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ idc_args2 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ idc_args3 <- ident.unify pattern.ident.Literal ##(projT2 args);
+ Some
+ (Base
+ (##(fancy.interp (invert_Some (to_fancy fancy_addm))
+ (let (x4, _) := idc_args1 in x4,
+ let (x4, _) := idc_args2 in x4,
+ let (x4, _) := idc_args3 in x4)%core))%expr)
+ else None
| None => None
end
| (@expr.Ident _ _ _ t idc @