aboutsummaryrefslogtreecommitdiff
path: root/src
diff options
context:
space:
mode:
Diffstat (limited to 'src')
-rw-r--r--src/AbstractInterpretation.v90
-rw-r--r--src/AbstractInterpretationProofs.v124
-rw-r--r--src/AbstractInterpretationWf.v90
-rw-r--r--src/BoundsPipeline.v1
-rw-r--r--src/Rewriter.v506
-rw-r--r--src/RewriterProofs.v20
-rw-r--r--src/RewriterRulesGood.v13
-rw-r--r--src/RewriterRulesInterpGood.v7
-rw-r--r--src/arith_with_casts_rewrite_head.out2280
-rw-r--r--src/fancy_with_casts_rewrite_head.out12579
-rw-r--r--src/strip_literal_casts_rewrite_head.out227
11 files changed, 6361 insertions, 9576 deletions
diff --git a/src/AbstractInterpretation.v b/src/AbstractInterpretation.v
index 4d82b869a..75ce71eff 100644
--- a/src/AbstractInterpretation.v
+++ b/src/AbstractInterpretation.v
@@ -732,11 +732,15 @@ Module Compilers.
end%under_lets.
(** We drop the state of higher-order arrows *)
- Fixpoint reify (is_let_bound : bool) {t} : value t -> type.for_each_lhs_of_arrow abstract_domain t -> UnderLets (@expr var t)
+ Fixpoint reify (annotate_with_state : bool) (is_let_bound : bool) {t} : value t -> type.for_each_lhs_of_arrow abstract_domain t -> UnderLets (@expr var t)
:= match t return value t -> type.for_each_lhs_of_arrow abstract_domain t -> UnderLets (@expr var t) with
| type.base t
=> fun '(st, v) 'tt
- => annotate is_let_bound t st v
+ => if annotate_with_state
+ then annotate is_let_bound t st v
+ else if is_let_bound
+ then UnderLets.UnderLet v (fun v' => UnderLets.Base ($v'))
+ else UnderLets.Base v
| type.arrow s d
=> fun f_e '(sv, dv)
=> let sv := match s with
@@ -745,10 +749,10 @@ Module Compilers.
end in
Base
(λ x , (UnderLets.to_expr
- (fx <-- f_e (@reflect _ (expr.Var x) sv);
- @reify false _ fx dv)))
+ (fx <-- f_e (@reflect annotate_with_state _ (expr.Var x) sv);
+ @reify annotate_with_state false _ fx dv)))
end%core%expr
- with reflect {t} : @expr var t -> abstract_domain t -> value t
+ with reflect (annotate_with_state : bool) {t} : @expr var t -> abstract_domain t -> value t
:= match t return @expr var t -> abstract_domain t -> value t with
| type.base t
=> fun e st => (st, e)
@@ -756,45 +760,45 @@ Module Compilers.
=> fun e absf
=> (fun v
=> let stv := state_of_value v in
- (rv <-- (@reify false s v bottom_for_each_lhs_of_arrow);
- Base (@reflect d (e @ rv) (absf stv))%expr))
+ (rv <-- (@reify annotate_with_state false s v bottom_for_each_lhs_of_arrow);
+ Base (@reflect annotate_with_state d (e @ rv) (absf stv))%expr))
end%under_lets.
- Fixpoint interp {t} (e : @expr value_with_lets t) : value_with_lets t
+ Fixpoint interp (annotate_with_state : bool) {t} (e : @expr value_with_lets t) : value_with_lets t
:= match e in expr.expr t return value_with_lets t with
| expr.Ident t idc => interp_ident _ idc (* Base (reflect (###idc) (abstract_interp_ident _ idc))*)
| expr.Var t v => v
- | expr.Abs s d f => Base (fun x => @interp d (f (Base x)))
+ | expr.Abs s d f => Base (fun x => @interp annotate_with_state d (f (Base x)))
| expr.App (type.base s) d f x
- => (x' <-- @interp _ x;
- f' <-- @interp (_ -> d)%etype f;
+ => (x' <-- @interp annotate_with_state _ x;
+ f' <-- @interp annotate_with_state (_ -> d)%etype f;
f' x')
| expr.App (type.arrow s' d') d f x
- => (x' <-- @interp (s' -> d')%etype x;
+ => (x' <-- @interp annotate_with_state (s' -> d')%etype x;
x'' <-- bottomify x';
- f' <-- @interp (_ -> d)%etype f;
+ f' <-- @interp annotate_with_state (_ -> d)%etype f;
f' x'')
| expr.LetIn (type.arrow _ _) B x f
- => (x' <-- @interp _ x;
- @interp _ (f (Base x')))
+ => (x' <-- @interp annotate_with_state _ x;
+ @interp annotate_with_state _ (f (Base x')))
| expr.LetIn (type.base A) B x f
- => (x' <-- @interp _ x;
- x'' <-- reify true (* this forces a let-binder here *) x' tt;
- @interp _ (f (Base (reflect x'' (state_of_value x')))))
+ => (x' <-- @interp annotate_with_state _ x;
+ x'' <-- reify annotate_with_state true (* this forces a let-binder here *) x' tt;
+ @interp annotate_with_state _ (f (Base (reflect annotate_with_state x'' (state_of_value x')))))
end%under_lets.
- Definition eval_with_bound' {t} (e : @expr value_with_lets t)
+ Definition eval_with_bound' (annotate_with_state : bool) {t} (e : @expr value_with_lets t)
(st : type.for_each_lhs_of_arrow abstract_domain t)
: expr t
- := UnderLets.to_expr (e' <-- interp e; reify false e' st).
+ := UnderLets.to_expr (e' <-- interp annotate_with_state e; reify annotate_with_state false e' st).
Definition eval' {t} (e : @expr value_with_lets t) : expr t
- := eval_with_bound' e bottom_for_each_lhs_of_arrow.
+ := eval_with_bound' false e bottom_for_each_lhs_of_arrow.
Definition eta_expand_with_bound' {t} (e : @expr var t)
(st : type.for_each_lhs_of_arrow abstract_domain t)
: expr t
- := UnderLets.to_expr (reify false (reflect e bottom) st).
+ := UnderLets.to_expr (reify true false (reflect true e bottom) st).
Section extract.
Context (ident_extract : forall t, ident t -> abstract_domain t).
@@ -832,7 +836,6 @@ Module Compilers.
Context (annotate_ident : forall t, abstract_domain' t -> option (ident (t -> t)))
(bottom' : forall A, abstract_domain' A)
(abstract_interp_ident : forall t, ident t -> type.interp abstract_domain' t)
- (update_literal_with_state : forall A : base.type.base, abstract_domain' A -> base.interp A -> base.interp A)
(extract_list_state : forall A, abstract_domain' (base.type.list A) -> option (list (abstract_domain' A)))
(is_annotated_for : forall t t', ident t -> abstract_domain' t' -> bool).
@@ -861,10 +864,7 @@ Module Compilers.
Definition annotate_base (is_let_bound : bool) {t : base.type.base}
(st : abstract_domain' t) (e : @expr var t)
: UnderLets (@expr var t)
- := match invert_Literal e with
- | Some v => Base ##(update_literal_with_state _ st v)
- | None => annotate_with_ident is_let_bound st e
- end%expr.
+ := annotate_with_ident is_let_bound st e.
Fixpoint annotate (is_let_bound : bool) {t : base.type} : abstract_domain' t -> @expr var t -> UnderLets (@expr var t)
:= match t return abstract_domain' t -> @expr var t -> UnderLets (@expr var t) with
@@ -904,10 +904,10 @@ Module Compilers.
Local Notation reflect := (@reflect base.type ident var abstract_domain' annotate bottom').
(** We manually rewrite with the rule for [nth_default], as the eliminator for eta-expanding lists in the input *)
- Definition interp_ident {t} (idc : ident t) : value_with_lets t
+ Definition interp_ident (annotate_with_state : bool) {t} (idc : ident t) : value_with_lets t
:= match idc in ident t return value_with_lets t with
| ident.List_nth_default T as idc
- => let default := reflect (###idc) (abstract_interp_ident _ idc) in
+ => let default := reflect annotate_with_state (###idc) (abstract_interp_ident _ idc) in
Base
(fun default_arg
=> default <-- default default_arg;
@@ -917,7 +917,7 @@ Module Compilers.
Base
(fun n_arg
=> default <-- default n_arg;
- ls' <-- @reify false (base.type.list T) ls_arg tt;
+ ls' <-- @reify annotate_with_state false (base.type.list T) ls_arg tt;
Base
(fst default,
match reflect_list ls', invert_Literal (snd n_arg) with
@@ -925,16 +925,16 @@ Module Compilers.
=> nth_default (snd default_arg) ls n
| _, _ => snd default
end))))
- | idc => Base (reflect (###idc) (abstract_interp_ident _ idc))
+ | idc => Base (reflect annotate_with_state (###idc) (abstract_interp_ident _ idc))
end%core%under_lets%expr.
- Definition eval_with_bound {t} (e : @expr value_with_lets t)
+ Definition eval_with_bound (annotate_with_state : bool) {t} (e : @expr value_with_lets t)
(st : type.for_each_lhs_of_arrow abstract_domain t)
: @expr var t
- := @eval_with_bound' base.type ident var abstract_domain' annotate bottom' (@interp_ident) t e st.
+ := @eval_with_bound' base.type ident var abstract_domain' annotate bottom' (@interp_ident annotate_with_state) annotate_with_state t e st.
Definition eval {t} (e : @expr value_with_lets t) : @expr var t
- := @eval' base.type ident var abstract_domain' annotate bottom' (@interp_ident) t e.
+ := @eval' base.type ident var abstract_domain' annotate bottom' (@interp_ident false) t e.
Definition eta_expand_with_bound {t} (e : @expr var t)
(st : type.for_each_lhs_of_arrow abstract_domain t)
@@ -999,32 +999,16 @@ Module Compilers.
:= ZRange.type.base.option.None.
Definition abstract_interp_ident t (idc : ident t) : type.interp abstract_domain' t
:= ZRange.ident.option.interp idc.
- Definition update_Z_literal_with_state : abstract_domain' base.type.Z -> Z -> Z
- := fun r n
- => match r with
- | Some r => if ZRange.type.base.is_bounded_by (t:=base.type.Z) r n
- then n
- else ident.cast_outside_of_range r n
- | None => n
- end.
- Definition update_literal_with_state (t : base.type.base) : abstract_domain' t -> base.interp t -> base.interp t
- := match t with
- | base.type.Z => update_Z_literal_with_state
- | base.type.unit
- | base.type.bool
- | base.type.nat
- => fun _ => id
- end.
Definition extract_list_state A (st : abstract_domain' (base.type.list A)) : option (list (abstract_domain' A))
:= st.
Definition eval_with_bound {var} {t} (e : @expr _ t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : expr t
:= (@partial.ident.eval_with_bound)
- var abstract_domain' annotate_ident bottom' abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for t e bound.
+ var abstract_domain' annotate_ident bottom' abstract_interp_ident extract_list_state is_annotated_for true t e bound.
Definition eta_expand_with_bound {var} {t} (e : @expr _ t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : expr t
:= (@partial.ident.eta_expand_with_bound)
- var abstract_domain' annotate_ident bottom' abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for t e bound.
+ var abstract_domain' annotate_ident bottom' abstract_interp_ident extract_list_state is_annotated_for t e bound.
Definition EvalWithBound {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : Expr t
:= fun var => eval_with_bound (e _) bound.
@@ -1034,7 +1018,7 @@ Module Compilers.
Definition eval {var} {t} (e : @expr _ t) : expr t
:= (@partial.ident.eval)
- var abstract_domain' (annotate_ident default_relax_zrange) bottom' abstract_interp_ident update_literal_with_state extract_list_state (is_annotated_for default_relax_zrange) t e.
+ var abstract_domain' (annotate_ident default_relax_zrange) bottom' abstract_interp_ident extract_list_state (is_annotated_for default_relax_zrange) t e.
Definition Eval {t} (e : Expr t) : Expr t
:= fun var => eval (e _).
Definition EtaExpandWithListInfoFromBound {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : Expr t
diff --git a/src/AbstractInterpretationProofs.v b/src/AbstractInterpretationProofs.v
index 6bf479dfe..c0f20474d 100644
--- a/src/AbstractInterpretationProofs.v
+++ b/src/AbstractInterpretationProofs.v
@@ -335,14 +335,29 @@ Module Compilers.
{ assumption. }
Qed.
+ Lemma related_bounded_value_Proper_eq {t}
+ : Proper (eq ==> eq ==> eq ==> Basics.impl) (@related_bounded_value t).
+ Proof using Type.
+ repeat intro; subst; assumption.
+ Qed.
+
+ Lemma related_bounded_value_Proper_interp_eq_base {t}
+ : Proper (eq ==> RelProd eq (fun x y => expr.interp ident_interp x = expr.interp ident_interp y) ==> eq ==> Basics.impl) (@related_bounded_value (type.base t)).
+ Proof using Type.
+ repeat intro; subst.
+ cbv [value RelProd relation_conjunction predicate_intersection pointwise_extension RelCompFun] in *.
+ destruct_head'_prod; destruct_head'_and; cbn [fst snd] in *; subst.
+ cbv [related_bounded_value] in *; destruct_head'_and; repeat apply conj; subst; (idtac + symmetry); assumption.
+ Qed.
+
Fixpoint interp_reify
- is_let_bound {t} st e v b_in
+ annotate_with_state is_let_bound {t} st e v b_in
(Hb : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) b_in)
(H : related_bounded_value st e v) {struct t}
: (forall arg1 arg2
(Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2)
(Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1),
- type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify is_let_bound t e b_in))) arg1
+ type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify annotate_with_state is_let_bound t e b_in))) arg1
= type.app_curried v arg2)
/\ (forall arg1
(Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1)
@@ -350,30 +365,31 @@ Module Compilers.
abstraction_relation'
_
(type.app_curried (fill_in_bottom_for_arrows st) b_in)
- (type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify is_let_bound t e b_in))) arg1))
+ (type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify annotate_with_state is_let_bound t e b_in))) arg1))
with interp_reflect
- {t} st e v
+ annotate_with_state {t} st e v
(Hst_Proper : Proper abstract_domain_R st)
(H_val : expr.interp ident_interp e == v)
(Hst1 : abstraction_relation (fill_in_bottom_for_arrows st) (expr.interp ident_interp e))
{struct t}
: related_bounded_value
st
- (@reflect t e st)
+ (@reflect annotate_with_state t e st)
v.
Proof using interp_annotate abstraction_relation'_Proper bottom'_related bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric.
all: destruct t as [t|s d];
[ clear interp_reify interp_reflect
- | pose proof (fun is_let_bound => interp_reify is_let_bound s) as interp_reify_s;
- pose proof (fun is_let_bound => interp_reify is_let_bound d) as interp_reify_d;
- pose proof (interp_reflect s) as interp_reflect_s;
- pose proof (interp_reflect d) as interp_reflect_d;
+ | pose proof (fun is_let_bound => interp_reify annotate_with_state is_let_bound s) as interp_reify_s;
+ pose proof (fun is_let_bound => interp_reify annotate_with_state is_let_bound d) as interp_reify_d;
+ pose proof (interp_reflect annotate_with_state s) as interp_reflect_s;
+ pose proof (interp_reflect annotate_with_state d) as interp_reflect_d;
clear interp_reify interp_reflect;
pose proof (@abstract_domain_R_bottom_fill_arrows s);
pose proof (@abstract_domain_R_bottom_fill_arrows d) ].
all: cbn [reify reflect] in *; fold (@reify) (@reflect) in *.
all: cbn [related_bounded_value type.related type.app_curried] in *.
all: cbn [UnderLets.interp expr.interp type.final_codomain type.andb_each_lhs_of_arrow type.is_base fst snd fill_in_bottom_for_arrows type.map_for_each_lhs_of_arrow type.for_each_lhs_of_arrow type.and_for_each_lhs_of_arrow partial.bottom_for_each_lhs_of_arrow partial.bottom] in *.
+ all: destruct annotate_with_state; try destruct is_let_bound.
all: repeat first [ reflexivity
| progress eta_expand
| progress cbv [type.is_not_higher_order] in *
@@ -408,7 +424,7 @@ Module Compilers.
try assumption; auto; []
| match goal with
| [ |- Proper ?R _ ] => (eapply PER_valid_l + eapply PER_valid_r); eassumption
- | [ |- @related_bounded_value ?t ?st1 (reflect _ ?st2) _ ]
+ | [ |- @related_bounded_value ?t ?st1 (reflect _ _ ?st2) _ ]
=> (tryif first [ constr_eq st1 st2 | has_evar st1 | has_evar st2 ]
then fail
else (eapply (@related_bounded_value_Proper1 t st2 st1);
@@ -427,7 +443,7 @@ Module Compilers.
end
| progress fold (@reify) (@reflect) (@type.interp) (@type.related) (@type.related_hetero) in *
| match goal with
- | [ |- type.related _ (expr.interp _ (UnderLets.interp _ (reify _ _ _))) _ ]
+ | [ |- type.related _ (expr.interp _ (UnderLets.interp _ (reify _ _ _ _))) _ ]
=> rewrite type.related_iff_app_curried
| [ |- type.related_hetero _ (@state_of_value ?t _) _ ]
=> is_var t; destruct t; cbv [state_of_value]; [ cbn | apply bottom_related ]
@@ -435,14 +451,14 @@ Module Compilers.
Qed.
Lemma interp_reify_first_order
- is_let_bound {t} st e v b_in
+ annotate_with_state is_let_bound {t} st e v b_in
(Ht : type.is_not_higher_order t = true)
(Hb : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) b_in)
(H : related_bounded_value st e v)
: (forall arg1 arg2
(Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2)
(Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1),
- type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify is_let_bound t e b_in))) arg1
+ type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify annotate_with_state is_let_bound t e b_in))) arg1
= type.app_curried v arg2)
/\ (forall arg1
(Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1)
@@ -450,21 +466,21 @@ Module Compilers.
abstraction_relation'
_
(type.app_curried st b_in)
- (type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify is_let_bound t e b_in))) arg1)).
+ (type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify annotate_with_state is_let_bound t e b_in))) arg1)).
Proof using interp_annotate abstraction_relation'_Proper bottom'_related bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric.
rewrite first_order_app_curried_fill_in_bottom_for_arrows_eq by assumption.
apply interp_reify; assumption.
Qed.
Lemma interp_reflect_first_order
- {t} st e v
+ annotate_with_state {t} st e v
(Ht : type.is_not_higher_order t = true)
(Hst_Proper : Proper abstract_domain_R st)
(H_val : expr.interp ident_interp e == v)
(Hst : abstraction_relation st (expr.interp ident_interp e))
: related_bounded_value
st
- (@reflect t e st)
+ (@reflect annotate_with_state t e st)
v.
Proof using interp_annotate abstraction_relation'_Proper bottom'_related bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric.
rewrite first_order_abstraction_relation_fill_in_bottom_for_arrows_iff in Hst by assumption.
@@ -507,7 +523,7 @@ Module Compilers.
related_bounded_value (ident_extract t idc) (UnderLets.interp ident_interp (interp_ident t idc)) (ident_interp t idc)).
Lemma interp_interp
- G G' {t} (e_st e1 e2 e3 : expr t)
+ annotate_with_state G G' {t} (e_st e1 e2 e3 : expr t)
(HG : forall t v1 v2 v3, List.In (existT _ t (v1, v2, v3)) G
-> related_bounded_value_with_lets v1 v2 v3)
(HG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G' -> v1 == v2)
@@ -515,13 +531,14 @@ Module Compilers.
(Hwf' : expr.wf G' e2 e3)
: related_bounded_value_with_lets
(extract' e_st)
- (interp e1)
+ (interp annotate_with_state e1)
(expr.interp (@ident_interp) e2).
Proof using interp_ident_Proper interp_annotate abstraction_relation'_Proper ident_interp_Proper' abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_Proper bottom'_related.
clear -ident_interp_Proper' interp_ident_Proper interp_annotate abstraction_relation'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_Proper bottom'_related HG HG' Hwf Hwf'.
cbv [related_bounded_value_with_lets] in *;
revert dependent G'; induction Hwf; intros;
- cbn [extract' interp expr.interp UnderLets.interp List.In related_bounded_value reify reflect] in *.
+ cbn [extract' interp expr.interp UnderLets.interp List.In related_bounded_value reify reflect] in *; cbv [Let_In] in *.
+ all: destruct annotate_with_state eqn:?.
all: repeat first [ progress intros
| progress subst
| progress inversion_sigma
@@ -538,8 +555,8 @@ Module Compilers.
| rewrite interp_annotate
| solve [ cbv [Proper respectful Basics.impl] in *; eauto using related_of_related_bounded_value, related_bounded_value_bottomify ]
| progress specialize_by_assumption
- | progress cbv [Let_In extract'] in *
- | progress cbn [state_of_value] in *
+ | progress cbv [Let_In] in *
+ | progress cbn [state_of_value extract'] in *
| progress expr.invert_subst
| match goal with
| [ |- abstract_domain ?t ] => exact (@bottom t)
@@ -553,6 +570,11 @@ Module Compilers.
| [ |- related_bounded_value _ (fst _, UnderLets.interp _ (annotate _ _ _ _)) _ ]
=> apply related_bounded_value_annotate_base
| [ H : context[match ?v with None => _ | _ => _ end] |- _ ] => destruct v eqn:?
+ | [ H : context[@related_bounded_value (type.base ?t) ?x _ ?y]
+ |- @related_bounded_value (type.base ?t) ?x _ ?y ]
+ => eapply related_bounded_value_Proper_interp_eq_base; [ reflexivity | split; hnf | reflexivity | eapply H ];
+ cbn [fst snd expr.interp];
+ [ reflexivity | reflexivity | .. ]
end
| apply conj
| match goal with
@@ -578,7 +600,7 @@ Module Compilers.
Qed.
Lemma interp_eval_with_bound'
- {t} (e_st e1 e2 : expr t)
+ annotate_with_state {t} (e_st e1 e2 : expr t)
(Hwf : expr.wf3 nil e_st e1 e2)
(Hwf' : expr.wf nil e2 e2)
(Ht : type.is_not_higher_order t = true)
@@ -587,7 +609,7 @@ Module Compilers.
: (forall arg1 arg2
(Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2)
(Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) st arg1),
- type.app_curried (expr.interp ident_interp (eval_with_bound' e1 st)) arg1
+ type.app_curried (expr.interp ident_interp (eval_with_bound' annotate_with_state e1 st)) arg1
= type.app_curried (expr.interp ident_interp e2) arg2)
/\ (forall arg1
(Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) st arg1)
@@ -595,7 +617,7 @@ Module Compilers.
abstraction_relation'
_
(extract_gen e_st st)
- (type.app_curried (expr.interp ident_interp (eval_with_bound' e1 st)) arg1)).
+ (type.app_curried (expr.interp ident_interp (eval_with_bound' annotate_with_state e1 st)) arg1)).
Proof using interp_annotate abstraction_relation'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_related interp_ident_Proper bottom'_Proper ident_interp_Proper'.
cbv [extract_gen extract' eval_with_bound'].
split; intros; rewrite UnderLets.interp_to_expr, UnderLets.interp_splice.
@@ -678,9 +700,6 @@ Module Compilers.
annotate_ident t st = Some idc
-> forall v, abstraction_relation' _ st v
-> ident_interp idc v = v)
- (interp_update_literal_with_state
- : forall (t : base.type.base) (st : abstract_domain' t) (v : base.interp t),
- abstraction_relation' t st v -> update_literal_with_state t st v = v)
(abstract_interp_ident_Proper'
: forall t idc, type.related_hetero (@abstraction_relation') (abstract_interp_ident t idc) (ident_interp idc))
(extract_list_state_related
@@ -697,9 +716,9 @@ Module Compilers.
Local Notation update_annotation := (@ident.update_annotation _ abstract_domain' annotate_ident is_annotated_for).
Local Notation annotate_with_ident := (@ident.annotate_with_ident _ abstract_domain' annotate_ident is_annotated_for).
- Local Notation annotate_base := (@ident.annotate_base _ abstract_domain' annotate_ident update_literal_with_state is_annotated_for).
- Local Notation annotate := (@ident.annotate _ abstract_domain' annotate_ident abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
- Local Notation interp_ident := (@ident.interp_ident _ abstract_domain' annotate_ident bottom' abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
+ Local Notation annotate_base := (@ident.annotate_base _ abstract_domain' annotate_ident is_annotated_for).
+ Local Notation annotate := (@ident.annotate _ abstract_domain' annotate_ident abstract_interp_ident extract_list_state is_annotated_for).
+ Local Notation interp_ident := (@ident.interp_ident _ abstract_domain' annotate_ident bottom' abstract_interp_ident extract_list_state is_annotated_for).
Local Notation related_bounded_value := (@related_bounded_value base.type ident abstract_domain' base.interp (@ident_interp) abstraction_relation' bottom' abstract_domain'_R).
Local Notation reify := (@reify base.type ident _ abstract_domain' annotate bottom').
Local Notation reflect := (@reflect base.type ident _ abstract_domain' annotate bottom').
@@ -746,9 +765,8 @@ Module Compilers.
(He : abstraction_relation' t st (expr.interp (t:=type.base (base.type.type_base t)) (@ident_interp) e))
: expr.interp (@ident_interp) (UnderLets.interp (@ident_interp) (@annotate_base is_let_bound t st e))
= expr.interp (@ident_interp) e.
- Proof using interp_annotate_ident interp_update_literal_with_state.
+ Proof using interp_annotate_ident.
cbv [annotate_base]; break_innermost_match; expr.invert_subst; cbv beta iota in *; subst.
- { cbn [expr.interp UnderLets.interp ident.smart_Literal ident_interp] in *; eauto. }
{ apply interp_annotate_with_ident; assumption. }
Qed.
@@ -756,7 +774,7 @@ Module Compilers.
(He : abstraction_relation' t st (expr.interp (t:=type.base t) (@ident_interp) e))
: expr.interp (@ident_interp) (UnderLets.interp (@ident_interp) (@annotate is_let_bound t st e))
= expr.interp (@ident_interp) e.
- Proof using interp_update_literal_with_state interp_annotate_ident abstract_interp_ident_Proper' extract_list_state_related extract_list_state_length_good bottom'_related.
+ Proof using interp_annotate_ident abstract_interp_ident_Proper' extract_list_state_related extract_list_state_length_good bottom'_related.
induction t; cbn [annotate]; auto using interp_annotate_base.
all: repeat first [ reflexivity
| progress subst
@@ -824,9 +842,9 @@ Module Compilers.
end ].
Qed.
- Lemma interp_ident_Proper_not_nth_default t idc
- : related_bounded_value (abstract_interp_ident t idc) (UnderLets.interp (@ident_interp) (Base (reflect (expr.Ident idc) (abstract_interp_ident _ idc)))) (ident_interp idc).
- Proof using abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_related extract_list_state_length_good extract_list_state_related interp_annotate_ident interp_update_literal_with_state abstract_interp_ident_Proper bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric.
+ Lemma interp_ident_Proper_not_nth_default annotate_with_state t idc
+ : related_bounded_value (abstract_interp_ident t idc) (UnderLets.interp (@ident_interp) (Base (reflect annotate_with_state (expr.Ident idc) (abstract_interp_ident _ idc)))) (ident_interp idc).
+ Proof using abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_related extract_list_state_length_good extract_list_state_related interp_annotate_ident abstract_interp_ident_Proper bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric.
cbn [UnderLets.interp].
eapply interp_reflect;
try first [ apply ident.gen_interp_Proper
@@ -836,9 +854,9 @@ Module Compilers.
eauto.
Qed.
- Lemma interp_ident_Proper_nth_default T (idc:=@ident.List_nth_default T)
- : related_bounded_value (abstract_interp_ident _ idc) (UnderLets.interp (@ident_interp) (interp_ident idc)) (ident_interp idc).
- Proof using abstract_interp_ident_Proper abstract_interp_ident_Proper' abstraction_relation'_Proper extract_list_state_length_good extract_list_state_related interp_annotate_ident interp_update_literal_with_state bottom'_related.
+ Lemma interp_ident_Proper_nth_default annotate_with_state T (idc:=@ident.List_nth_default T)
+ : related_bounded_value (abstract_interp_ident _ idc) (UnderLets.interp (@ident_interp) (interp_ident annotate_with_state idc)) (ident_interp idc).
+ Proof using abstract_interp_ident_Proper abstract_interp_ident_Proper' abstraction_relation'_Proper extract_list_state_length_good extract_list_state_related interp_annotate_ident bottom'_related.
subst idc; cbn [interp_ident reify reflect fst snd UnderLets.interp ident_interp related_bounded_value abstract_domain value].
cbv [abstract_domain]; cbn [type.interp bottom_for_each_lhs_of_arrow state_of_value fst snd].
repeat first [ progress intros
@@ -858,6 +876,7 @@ Module Compilers.
| apply (@abstract_interp_ident_Proper'' _ (@ident.List_nth_default T))
| apply conj
| rewrite map_nth_default_always
+ | rewrite expr.interp_reify_list
| match goal with
| [ H : context[expr.interp _ (UnderLets.interp _ (annotate _ _ _))] |- _ ]
=> rewrite interp_annotate in H
@@ -869,20 +888,20 @@ Module Compilers.
end ].
Qed.
- Lemma interp_ident_Proper t idc
- : related_bounded_value (abstract_interp_ident t idc) (UnderLets.interp (@ident_interp) (interp_ident idc)) (ident_interp idc).
+ Lemma interp_ident_Proper annotate_with_state t idc
+ : related_bounded_value (abstract_interp_ident t idc) (UnderLets.interp (@ident_interp) (interp_ident annotate_with_state idc)) (ident_interp idc).
Proof.
pose idc as idc'.
- destruct idc; first [ refine (@interp_ident_Proper_not_nth_default _ idc')
- | refine (@interp_ident_Proper_nth_default _) ].
+ destruct idc; first [ refine (@interp_ident_Proper_not_nth_default _ _ idc')
+ | refine (@interp_ident_Proper_nth_default _ _) ].
Qed.
- Local Notation eval_with_bound := (@partial.ident.eval_with_bound _ abstract_domain' annotate_ident bottom' abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
- Local Notation eta_expand_with_bound := (@partial.ident.eta_expand_with_bound _ abstract_domain' annotate_ident bottom' abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
+ Local Notation eval_with_bound := (@partial.ident.eval_with_bound _ abstract_domain' annotate_ident bottom' abstract_interp_ident extract_list_state is_annotated_for).
+ Local Notation eta_expand_with_bound := (@partial.ident.eta_expand_with_bound _ abstract_domain' annotate_ident bottom' abstract_interp_ident extract_list_state is_annotated_for).
Local Notation extract := (@ident.extract abstract_domain' bottom' abstract_interp_ident).
Lemma interp_eval_with_bound
- {t} (e_st e1 e2 : expr t)
+ annotate_with_state {t} (e_st e1 e2 : expr t)
(Hwf : expr.wf3 nil e_st e1 e2)
(Hwf' : expr.wf nil e2 e2)
(Ht : type.is_not_higher_order t = true)
@@ -891,7 +910,7 @@ Module Compilers.
: (forall arg1 arg2
(Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2)
(Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) st arg1),
- type.app_curried (expr.interp (@ident_interp) (eval_with_bound e1 st)) arg1
+ type.app_curried (expr.interp (@ident_interp) (eval_with_bound annotate_with_state e1 st)) arg1
= type.app_curried (expr.interp (@ident_interp) e2) arg2)
/\ (forall arg1
(Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) st arg1)
@@ -899,7 +918,7 @@ Module Compilers.
abstraction_relation'
_
(extract e_st st)
- (type.app_curried (expr.interp (@ident_interp) (eval_with_bound e1 st)) arg1)).
+ (type.app_curried (expr.interp (@ident_interp) (eval_with_bound annotate_with_state e1 st)) arg1)).
Proof. cbv [extract eval_with_bound]; apply @interp_eval_with_bound' with (abstract_domain'_R:=abstract_domain'_R); auto using interp_annotate, interp_ident_Proper, ident.gen_interp_Proper. Qed.
Lemma interp_eta_expand_with_bound
@@ -951,13 +970,6 @@ Module Compilers.
: type.related_hetero (@abstraction_relation') (@abstract_interp_ident t idc) (@ident.gen_interp cast_outside_of_range _ idc).
Proof using Type. apply ZRange.ident.option.interp_related. Qed.
- Lemma interp_update_literal_with_state {t : base.type.base} st v
- : @abstraction_relation' t st v -> @update_literal_with_state t st v = v.
- Proof using Type.
- cbv [abstraction_relation' update_literal_with_state update_Z_literal_with_state ZRange.type.base.option.is_bounded_by];
- break_innermost_match; try congruence; reflexivity.
- Qed.
-
Lemma extract_list_state_related {t} st v ls
: @abstraction_relation' _ st v
-> @extract_list_state t st = Some ls
@@ -1034,7 +1046,7 @@ Module Compilers.
: @annotate_ident relax_zrange t st1 = @annotate_ident relax_zrange t st2.
Proof using Type. congruence. Qed.
- Local Hint Resolve interp_annotate_ident interp_update_literal_with_state abstract_interp_ident_related.
+ Local Hint Resolve interp_annotate_ident abstract_interp_ident_related.
Lemma interp_eval_with_bound
cast_outside_of_range
diff --git a/src/AbstractInterpretationWf.v b/src/AbstractInterpretationWf.v
index a2ba63c01..a6a2d8d6b 100644
--- a/src/AbstractInterpretationWf.v
+++ b/src/AbstractInterpretationWf.v
@@ -171,23 +171,23 @@ Module Compilers.
| eapply UnderLets.wf_Proper_list; [ | | solve [ eauto ] ] ].
Qed.
- Fixpoint wf_reify (is_let_bound : bool) G {t}
+ Fixpoint wf_reify (annotate_with_state : bool) (is_let_bound : bool) G {t}
: forall v1 v2 (Hv : @wf_value G t v1 v2)
s1 s2 (Hs : type.and_for_each_lhs_of_arrow (@abstract_domain_R) s1 s2),
- UnderLets.wf (fun G' => expr.wf G') G (@reify1 is_let_bound t v1 s1) (@reify2 is_let_bound t v2 s2)
- with wf_reflect G {t}
+ UnderLets.wf (fun G' => expr.wf G') G (@reify1 annotate_with_state is_let_bound t v1 s1) (@reify2 annotate_with_state is_let_bound t v2 s2)
+ with wf_reflect (annotate_with_state : bool) G {t}
: forall e1 e2 (He : expr.wf G e1 e2)
s1 s2 (Hs : abstract_domain_R s1 s2),
- @wf_value G t (@reflect1 t e1 s1) (@reflect2 t e2 s2).
+ @wf_value G t (@reflect1 annotate_with_state t e1 s1) (@reflect2 annotate_with_state t e2 s2).
Proof using annotate_Proper bottom'_Proper.
all: clear -wf_reflect wf_reify annotate_Proper type_base bottom'_Proper.
all: pose proof (@bottom_for_each_lhs_of_arrow_Proper); cbv [Proper abstract_domain_R] in *.
all: destruct t as [t|s d];
[ clear wf_reify wf_reflect
- | pose proof (fun G => wf_reflect G s) as wf_reflect_s;
- pose proof (fun G => wf_reflect G d) as wf_reflect_d;
- pose proof (fun G => wf_reify false G s) as wf_reify_s;
- pose proof (fun G => wf_reify false G d) as wf_reify_d;
+ | pose proof (fun G => wf_reflect annotate_with_state G s) as wf_reflect_s;
+ pose proof (fun G => wf_reflect annotate_with_state G d) as wf_reflect_d;
+ pose proof (fun G => wf_reify annotate_with_state false G s) as wf_reify_s;
+ pose proof (fun G => wf_reify annotate_with_state false G d) as wf_reify_d;
pose proof (@bottom_Proper s);
clear wf_reify wf_reflect ].
all: cbn [reify reflect] in *.
@@ -208,9 +208,9 @@ Module Compilers.
| [ |- expr.wf _ _ _ ] => constructor
| [ He : forall seg G' sv1 sv2, G' = (seg ++ ?G)%list -> _ |- UnderLets.wf _ (?v :: ?G) (UnderLets.splice _ _) (UnderLets.splice _ _) ]
=> eapply UnderLets.wf_splice; [ apply (He (cons v nil)) | ]
- | [ |- UnderLets.wf _ _ (UnderLets.splice (reify1 _ _ _) _) (UnderLets.splice (reify2 _ _ _) _) ]
+ | [ |- UnderLets.wf _ _ (UnderLets.splice (reify1 _ _ _ _) _) (UnderLets.splice (reify2 _ _ _ _) _) ]
=> eapply UnderLets.wf_splice; [ apply wf_reify_s || apply wf_reify_d | ]
- | [ |- wf_value _ (reflect1 _ _) (reflect2 _ _) ] => apply wf_reflect_s || apply wf_reflect_d
+ | [ |- wf_value _ (reflect1 _ _ _) (reflect2 _ _ _) ] => apply wf_reflect_s || apply wf_reflect_d
| [ H : wf_value _ ?x ?y |- wf_value _ ?x ?y ]
=> eapply wf_value_Proper_list; [ | eassumption ]
| [ H : forall x y, ?R x y -> ?R' (?f x) (?g y) |- ?R' (?f _) (?g _) ]
@@ -259,10 +259,10 @@ Module Compilers.
end
| break_innermost_match_step ].
- Lemma wf_interp G G' {t} (e1 : @expr (@value_with_lets1) t) (e2 : @expr (@value_with_lets2) t)
+ Lemma wf_interp (annotate_with_state : bool) G G' {t} (e1 : @expr (@value_with_lets1) t) (e2 : @expr (@value_with_lets2) t)
(Hwf : expr.wf G e1 e2)
(HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2)
- : wf_value_with_lets G' (interp1 e1) (interp2 e2).
+ : wf_value_with_lets G' (interp1 annotate_with_state e1) (interp2 annotate_with_state e2).
Proof using annotate_Proper bottom'_Proper interp_ident_Proper.
revert dependent G'; induction Hwf; intros; cbn [interp];
try solve [ apply interp_ident_Proper; auto
@@ -270,9 +270,9 @@ Module Compilers.
wf_interp_t.
Qed.
- Lemma wf_eval_with_bound' G G' {t} e1 e2 (He : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2)
+ Lemma wf_eval_with_bound' (annotate_with_state : bool) G G' {t} e1 e2 (He : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2)
(HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2)
- : expr.wf G' (@eval_with_bound'1 t e1 st1) (@eval_with_bound'2 t e2 st2).
+ : expr.wf G' (@eval_with_bound'1 annotate_with_state t e1 st1) (@eval_with_bound'2 annotate_with_state t e2 st2).
Proof using annotate_Proper bottom'_Proper interp_ident_Proper.
eapply UnderLets.wf_to_expr, UnderLets.wf_splice.
{ eapply wf_interp; solve [ eauto ]. }
@@ -303,7 +303,6 @@ Module Compilers.
Context (annotate_ident : forall t, abstract_domain' t -> option (ident (t -> t)))
(bottom' : forall A, abstract_domain' A)
(abstract_interp_ident : forall t, ident t -> type.interp abstract_domain' t)
- (update_literal_with_state : forall A : base.type.base, abstract_domain' A -> base.interp A -> base.interp A)
(extract_list_state : forall A, abstract_domain' (base.type.list A) -> option (list (abstract_domain' A)))
(is_annotated_for : forall t t', ident t -> abstract_domain' t' -> bool).
Context (abstract_domain'_R : forall t, abstract_domain' t -> abstract_domain' t -> Prop).
@@ -311,7 +310,6 @@ Module Compilers.
Context {annotate_ident_Proper : forall t, Proper (abstract_domain'_R t ==> eq) (annotate_ident t)}
{abstract_interp_ident_Proper : forall t, Proper (eq ==> @abstract_domain_R t) (abstract_interp_ident t)}
{bottom'_Proper : forall t, Proper (abstract_domain'_R t) (bottom' t)}
- {update_literal_with_state_Proper : forall t, Proper (abstract_domain'_R (base.type.type_base t) ==> eq ==> eq) (update_literal_with_state t)}
{is_annotated_for_Proper : forall t t', Proper (eq ==> abstract_domain'_R _ ==> eq) (@is_annotated_for t t')}
(extract_list_state_length : forall t v1 v2, abstract_domain'_R _ v1 v2 -> option_map (@length _) (extract_list_state t v1) = option_map (@length _) (extract_list_state t v2))
(extract_list_state_rel : forall t v1 v2, abstract_domain'_R _ v1 v2 -> forall l1 l2, extract_list_state t v1 = Some l1 -> extract_list_state t v2 = Some l2 -> forall vv1 vv2, List.In (vv1, vv2) (List.combine l1 l2) -> @abstract_domain'_R t vv1 vv2).
@@ -325,14 +323,14 @@ Module Compilers.
Local Notation update_annotation1 := (@ident.update_annotation var1 abstract_domain' annotate_ident is_annotated_for).
Local Notation update_annotation2 := (@ident.update_annotation var2 abstract_domain' annotate_ident is_annotated_for).
- Local Notation annotate1 := (@ident.annotate var1 abstract_domain' annotate_ident abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
- Local Notation annotate2 := (@ident.annotate var2 abstract_domain' annotate_ident abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
- Local Notation annotate_base1 := (@ident.annotate_base var1 abstract_domain' annotate_ident update_literal_with_state is_annotated_for).
- Local Notation annotate_base2 := (@ident.annotate_base var2 abstract_domain' annotate_ident update_literal_with_state is_annotated_for).
+ Local Notation annotate1 := (@ident.annotate var1 abstract_domain' annotate_ident abstract_interp_ident extract_list_state is_annotated_for).
+ Local Notation annotate2 := (@ident.annotate var2 abstract_domain' annotate_ident abstract_interp_ident extract_list_state is_annotated_for).
+ Local Notation annotate_base1 := (@ident.annotate_base var1 abstract_domain' annotate_ident is_annotated_for).
+ Local Notation annotate_base2 := (@ident.annotate_base var2 abstract_domain' annotate_ident is_annotated_for).
Local Notation annotate_with_ident1 := (@ident.annotate_with_ident var1 abstract_domain' annotate_ident is_annotated_for).
Local Notation annotate_with_ident2 := (@ident.annotate_with_ident var2 abstract_domain' annotate_ident is_annotated_for).
- Local Notation interp_ident1 := (@ident.interp_ident var1 abstract_domain' annotate_ident bottom' abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
- Local Notation interp_ident2 := (@ident.interp_ident var2 abstract_domain' annotate_ident bottom' abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
+ Local Notation interp_ident1 := (@ident.interp_ident var1 abstract_domain' annotate_ident bottom' abstract_interp_ident extract_list_state is_annotated_for).
+ Local Notation interp_ident2 := (@ident.interp_ident var2 abstract_domain' annotate_ident bottom' abstract_interp_ident extract_list_state is_annotated_for).
Local Notation reflect1 := (@reflect base.type ident var1 abstract_domain' annotate1 bottom').
Local Notation reflect2 := (@reflect base.type ident var2 abstract_domain' annotate2 bottom').
@@ -382,7 +380,7 @@ Module Compilers.
v1 v2 (Hv : abstract_domain'_R t v1 v2)
e1 e2 (He : expr.wf G e1 e2)
: UnderLets.wf (fun G' => expr.wf G') G (@annotate_base1 is_let_bound t v1 e1) (@annotate_base2 is_let_bound t v2 e2).
- Proof using abstract_interp_ident_Proper annotate_ident_Proper update_literal_with_state_Proper is_annotated_for_Proper.
+ Proof using abstract_interp_ident_Proper annotate_ident_Proper is_annotated_for_Proper.
cbv [ident.annotate_base];
repeat first [ apply wf_annotate_with_ident
| break_innermost_match_step
@@ -409,7 +407,7 @@ Module Compilers.
v1 v2 (Hv : abstract_domain'_R t v1 v2)
e1 e2 (He : expr.wf G e1 e2)
: UnderLets.wf (fun G' => expr.wf G') G (@annotate1 is_let_bound t v1 e1) (@annotate2 is_let_bound t v2 e2).
- Proof using abstract_interp_ident_Proper annotate_ident_Proper extract_list_state_length extract_list_state_rel update_literal_with_state_Proper is_annotated_for_Proper.
+ Proof using abstract_interp_ident_Proper annotate_ident_Proper extract_list_state_length extract_list_state_rel is_annotated_for_Proper.
revert dependent G; induction t; intros;
cbn [ident.annotate]; try apply wf_annotate_base; trivial.
all: repeat first [ lazymatch goal with
@@ -491,9 +489,9 @@ Module Compilers.
let b' := type_of_value b in
constr:(type.arrow a' b')
end.
- Lemma wf_interp_ident_nth_default G T
- : wf_value_with_lets G (@interp_ident1 _ (@ident.List_nth_default T)) (@interp_ident2 _ (@ident.List_nth_default T)).
- Proof using abstract_interp_ident_Proper annotate_ident_Proper extract_list_state_length extract_list_state_rel update_literal_with_state_Proper is_annotated_for_Proper.
+ Lemma wf_interp_ident_nth_default (annotate_with_state : bool) G T
+ : wf_value_with_lets G (@interp_ident1 annotate_with_state _ (@ident.List_nth_default T)) (@interp_ident2 annotate_with_state _ (@ident.List_nth_default T)).
+ Proof using abstract_interp_ident_Proper annotate_ident_Proper extract_list_state_length extract_list_state_rel is_annotated_for_Proper.
cbv [wf_value_with_lets wf_value ident.interp_ident]; constructor; cbn -[abstract_domain_R abstract_domain].
{ intros; subst.
destruct_head'_prod; destruct_head'_and; cbn [fst snd] in *.
@@ -599,9 +597,9 @@ Module Compilers.
end ]. }
Qed.
- Lemma wf_interp_ident_not_nth_default G {t} (idc : ident t)
- : wf_value_with_lets G (Base (reflect1 (###idc)%expr (abstract_interp_ident _ idc))) (Base (reflect2 (###idc)%expr (abstract_interp_ident _ idc))).
- Proof using abstract_interp_ident_Proper annotate_ident_Proper bottom'_Proper extract_list_state_length extract_list_state_rel update_literal_with_state_Proper is_annotated_for_Proper.
+ Lemma wf_interp_ident_not_nth_default (annotate_with_state : bool) G {t} (idc : ident t)
+ : wf_value_with_lets G (Base (reflect1 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc))) (Base (reflect2 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc))).
+ Proof using abstract_interp_ident_Proper annotate_ident_Proper bottom'_Proper extract_list_state_length extract_list_state_rel is_annotated_for_Proper.
constructor; eapply wf_reflect;
solve [ apply bottom'_Proper
| apply wf_annotate
@@ -609,43 +607,43 @@ Module Compilers.
| apply abstract_interp_ident_Proper; reflexivity ].
Qed.
- Lemma wf_interp_ident G {t} idc1 idc2 (Hidc : idc1 = idc2)
- : wf_value_with_lets G (@interp_ident1 t idc1) (@interp_ident2 t idc2).
- Proof using abstract_interp_ident_Proper annotate_ident_Proper bottom'_Proper extract_list_state_length extract_list_state_rel update_literal_with_state_Proper is_annotated_for_Proper.
+ Lemma wf_interp_ident (annotate_with_state : bool) G {t} idc1 idc2 (Hidc : idc1 = idc2)
+ : wf_value_with_lets G (@interp_ident1 annotate_with_state t idc1) (@interp_ident2 annotate_with_state t idc2).
+ Proof using abstract_interp_ident_Proper annotate_ident_Proper bottom'_Proper extract_list_state_length extract_list_state_rel is_annotated_for_Proper.
cbv [wf_value_with_lets ident.interp_ident]; subst idc2; destruct idc1;
first [ apply wf_interp_ident_nth_default
| apply wf_interp_ident_not_nth_default ].
Qed.
- Local Notation eval_with_bound1 := (@partial.ident.eval_with_bound var1 abstract_domain' annotate_ident bottom' abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
- Local Notation eval_with_bound2 := (@partial.ident.eval_with_bound var2 abstract_domain' annotate_ident bottom' abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
- Lemma wf_eval_with_bound {t} G G' e1 e2 (Hwf : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2)
+ Local Notation eval_with_bound1 := (@partial.ident.eval_with_bound var1 abstract_domain' annotate_ident bottom' abstract_interp_ident extract_list_state is_annotated_for).
+ Local Notation eval_with_bound2 := (@partial.ident.eval_with_bound var2 abstract_domain' annotate_ident bottom' abstract_interp_ident extract_list_state is_annotated_for).
+ Lemma wf_eval_with_bound (annotate_with_state : bool) {t} G G' e1 e2 (Hwf : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2)
(HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2)
- : expr.wf G' (@eval_with_bound1 t e1 st1) (@eval_with_bound2 t e2 st2).
- Proof using abstract_interp_ident_Proper annotate_ident_Proper bottom'_Proper extract_list_state_length extract_list_state_rel update_literal_with_state_Proper is_annotated_for_Proper.
+ : expr.wf G' (@eval_with_bound1 annotate_with_state t e1 st1) (@eval_with_bound2 annotate_with_state t e2 st2).
+ Proof using abstract_interp_ident_Proper annotate_ident_Proper bottom'_Proper extract_list_state_length extract_list_state_rel is_annotated_for_Proper.
eapply wf_eval_with_bound';
solve [ eassumption
| eapply wf_annotate
| eapply wf_interp_ident ].
Qed.
- Local Notation eval1 := (@partial.ident.eval var1 abstract_domain' annotate_ident bottom' abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
- Local Notation eval2 := (@partial.ident.eval var2 abstract_domain' annotate_ident bottom' abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
+ Local Notation eval1 := (@partial.ident.eval var1 abstract_domain' annotate_ident bottom' abstract_interp_ident extract_list_state is_annotated_for).
+ Local Notation eval2 := (@partial.ident.eval var2 abstract_domain' annotate_ident bottom' abstract_interp_ident extract_list_state is_annotated_for).
Lemma wf_eval {t} G G' e1 e2 (Hwf : expr.wf G e1 e2)
(HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2)
: expr.wf G' (@eval1 t e1) (@eval2 t e2).
- Proof using abstract_interp_ident_Proper annotate_ident_Proper bottom'_Proper extract_list_state_length extract_list_state_rel update_literal_with_state_Proper is_annotated_for_Proper.
+ Proof using abstract_interp_ident_Proper annotate_ident_Proper bottom'_Proper extract_list_state_length extract_list_state_rel is_annotated_for_Proper.
eapply wf_eval';
solve [ eassumption
| eapply wf_annotate
| eapply wf_interp_ident ].
Qed.
- Local Notation eta_expand_with_bound1 := (@partial.ident.eta_expand_with_bound var1 abstract_domain' annotate_ident bottom' abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
- Local Notation eta_expand_with_bound2 := (@partial.ident.eta_expand_with_bound var2 abstract_domain' annotate_ident bottom' abstract_interp_ident update_literal_with_state extract_list_state is_annotated_for).
+ Local Notation eta_expand_with_bound1 := (@partial.ident.eta_expand_with_bound var1 abstract_domain' annotate_ident bottom' abstract_interp_ident extract_list_state is_annotated_for).
+ Local Notation eta_expand_with_bound2 := (@partial.ident.eta_expand_with_bound var2 abstract_domain' annotate_ident bottom' abstract_interp_ident extract_list_state is_annotated_for).
Lemma wf_eta_expand_with_bound {t} G e1 e2 (Hwf : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2)
: expr.wf G (@eta_expand_with_bound1 t e1 st1) (@eta_expand_with_bound2 t e2 st2).
- Proof using abstract_interp_ident_Proper annotate_ident_Proper bottom'_Proper extract_list_state_length extract_list_state_rel update_literal_with_state_Proper is_annotated_for_Proper.
+ Proof using abstract_interp_ident_Proper annotate_ident_Proper bottom'_Proper extract_list_state_length extract_list_state_rel is_annotated_for_Proper.
eapply wf_eta_expand_with_bound';
solve [ eassumption
| eapply wf_annotate
@@ -697,10 +695,6 @@ Module Compilers.
end ].
Qed.
- Global Instance update_literal_with_state_Proper {t}
- : Proper (abstract_domain'_R (base.type.type_base t) ==> eq ==> eq) (update_literal_with_state t).
- Proof. repeat intro; congruence. Qed.
-
Global Instance extract_list_state_Proper {t}
: Proper (abstract_domain'_R _ ==> option_eq (SetoidList.eqlistA (@abstract_domain'_R t)))
(extract_list_state t).
diff --git a/src/BoundsPipeline.v b/src/BoundsPipeline.v
index 0299f2e9e..1c6cec4c7 100644
--- a/src/BoundsPipeline.v
+++ b/src/BoundsPipeline.v
@@ -311,6 +311,7 @@ Module Pipeline.
=> RewriteRules.RewriteToFancyWithCasts invert_low invert_high value_range flag_range E
| None => E
end in
+ let E := RewriteRules.RewriteStripLiteralCasts E in
Success E
| inr (inl (b, E))
=> Error (Computed_bounds_are_not_tight_enough b out_bounds E arg_bounds)
diff --git a/src/Rewriter.v b/src/Rewriter.v
index 5f10869b2..a79bcb18f 100644
--- a/src/Rewriter.v
+++ b/src/Rewriter.v
@@ -359,6 +359,7 @@ Module Compilers.
Notation "x 'mod' y" := (#ident.Z_modulo @ x @ y)%pattern : pattern_scope.
Notation "- x" := (#ident.Z_opp @ x) : pattern_scope.
+ Notation "#?ℤ'" := (#ident.Z_cast @ #?ℤ) : pattern_scope.
Notation "??'" := (#ident.Z_cast @ Wildcard _) : pattern_scope.
Notation "x -' y" := (#ident.Z_cast @ (#ident.Z_sub @ x @ y)) : pattern_scope.
Notation "x +' y" := (#ident.Z_cast @ (#ident.Z_add @ x @ y)) : pattern_scope.
@@ -1747,34 +1748,34 @@ Module Compilers.
Definition arith_with_casts_rewrite_rules : rewrite_rulesT
:= [make_rewrite (#(@pattern.ident.fst '1 '2) @ (??, ??)) (fun _ _ x y => x)
- ; make_rewrite (#(@pattern.ident.snd '1 '2) @ (??, ??)) (fun _ x _ y => y)
+ ; make_rewrite (#(@pattern.ident.snd '1 '2) @ (??, ??)) (fun _ _ x y => y)
- ; make_rewriteo (??') (fun r v => ##(lower r) when lower r =? upper r)
+ ; make_rewriteo (??') (fun r v => cst r (##(lower r)) when lower r =? upper r)
; make_rewriteo
- (#?ℤ +' ??')
- (fun rp z rv v => cst rv v when (z =? 0) && (ZRange.normalize rv <=? ZRange.normalize rp)%zrange)
+ (#?ℤ' + ?? )
+ (fun rz z v => v when (z =? 0) && (is_bounded_by_bool z (ZRange.normalize rz)))
; make_rewriteo
- (??' +' #?ℤ )
- (fun rp rv v z => cst rv v when (z =? 0) && (ZRange.normalize rv <=? ZRange.normalize rp)%zrange)
+ (?? + #?ℤ')
+ (fun v rz z => v when (z =? 0) && (is_bounded_by_bool z (ZRange.normalize rz)))
; make_rewriteo
- (#?ℤ - (-'??'))
- (fun z rnv rv v => cst rv v when (z =? 0) && (ZRange.normalize rv <=? -ZRange.normalize rnv)%zrange)
- ; make_rewriteo (#?ℤ - ?? ) (fun z v => -v when z =? 0)
+ (#?ℤ' - (-'??'))
+ (fun rz z rnv rv v => cst rv v when (z =? 0) && (ZRange.normalize rv <=? -ZRange.normalize rnv)%zrange && (is_bounded_by_bool z rz))
+ ; make_rewriteo (#?ℤ' - ?? ) (fun rz z v => -v when (z =? 0) && is_bounded_by_bool z (ZRange.normalize rz))
- ; make_rewriteo (#?ℤ << ??) (fun x y => ##0 when x =? 0)
+ ; make_rewriteo (#?ℤ' << ??) (fun rx x y => ##0 when (x =? 0) && is_bounded_by_bool x (ZRange.normalize rx))
; make_rewriteo (-(-'??')) (fun rnv rv v => cst rv v when (ZRange.normalize rv <=? -ZRange.normalize rnv)%zrange)
- ; make_rewriteo (#pattern.ident.Z_mul_split @ #?ℤ @ #?ℤ @ ??) (fun s xx y => (##0, ##0)%Z when xx =? 0)
- ; make_rewriteo (#pattern.ident.Z_mul_split @ #?ℤ @ ?? @ #?ℤ) (fun s y xx => (##0, ##0)%Z when xx =? 0)
+ ; make_rewriteo (#pattern.ident.Z_mul_split @ ?? @ #?ℤ' @ ??) (fun s rxx xx y => (cst r[0~>0] ##0, cst r[0~>0] ##0)%Z when (xx =? 0) && is_bounded_by_bool xx (ZRange.normalize rxx))
+ ; make_rewriteo (#pattern.ident.Z_mul_split @ ?? @ ?? @ #?ℤ') (fun s y rxx xx => (cst r[0~>0] ##0, cst r[0~>0] ##0)%Z when (xx =? 0) && is_bounded_by_bool xx (ZRange.normalize rxx))
; make_rewriteo
- (#pattern.ident.Z_mul_split @ #?ℤ @ #?ℤ @ ??')
- (fun s xx ry y => (cst ry y, ##0)%Z when (xx =? 1) && (ZRange.normalize ry <=? r[0~>s-1])%zrange)
+ (#pattern.ident.Z_mul_split @ #?ℤ' @ #?ℤ' @ ??')
+ (fun rs s rxx xx ry y => (cst ry y, cst r[0~>0] ##0)%Z when (xx =? 1) && (ZRange.normalize ry <=? r[0~>s-1])%zrange && is_bounded_by_bool s (ZRange.normalize rs) && is_bounded_by_bool xx (ZRange.normalize rxx))
; make_rewriteo
- (#pattern.ident.Z_mul_split @ #?ℤ @ ??' @ #?ℤ)
- (fun s ry y xx => (cst ry y, ##0)%Z when (xx =? 1) && (ZRange.normalize ry <=? r[0~>s-1])%zrange)
+ (#pattern.ident.Z_mul_split @ #?ℤ' @ ??' @ #?ℤ')
+ (fun rs s ry y rxx xx => (cst ry y, cst r[0~>0] ##0)%Z when (xx =? 1) && (ZRange.normalize ry <=? r[0~>s-1])%zrange && is_bounded_by_bool s (ZRange.normalize rs) && is_bounded_by_bool xx (ZRange.normalize rxx))
(*
; make_rewriteo
(#pattern.ident.Z_mul_split @ #?ℤ @ #?ℤ @ ??')
@@ -1796,14 +1797,14 @@ Module Compilers.
=> (llet2_opp2 rvc (#ident.Z_sub_get_borrow @ s @ x @ cst ry y))
when (ZRange.normalize ry <=? -ZRange.normalize rny)%zrange)
; make_rewriteol
- (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_get_carry @ ?? @ #?ℤ @ ??))
- (fun rvc s yy x
- => (llet2_opp2 rvc (#ident.Z_sub_get_borrow @ s @ x @ ##(-yy)%Z))
- when yy <? 0)
+ (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_get_carry @ ?? @ #?ℤ' @ ??))
+ (fun rvc s ryy yy x
+ => (llet2_opp2 rvc (#ident.Z_sub_get_borrow @ s @ x @ cst (ZRange.opp ryy) ##(-yy)%Z))
+ when (yy <? 0) && is_bounded_by_bool yy (ZRange.normalize ryy))
; make_rewriteol
- (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_get_carry @ ?? @ ?? @ #?ℤ))
- (fun rvc s x yy => (llet2_opp2 rvc (#ident.Z_sub_get_borrow @ s @ x @ ##(-yy)%Z))
- when yy <? 0)
+ (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_get_carry @ ?? @ ?? @ #?ℤ'))
+ (fun rvc s x ryy yy => (llet2_opp2 rvc (#ident.Z_sub_get_borrow @ s @ x @ cst (ZRange.opp ryy) ##(-yy)%Z))
+ when (yy <? 0) && is_bounded_by_bool yy (ZRange.normalize ryy))
; make_rewriteol
@@ -1817,82 +1818,82 @@ Module Compilers.
=> (llet2_opp2 rvc (#ident.Z_sub_with_get_borrow @ s @ (cst rc c) @ x @ (cst ry y)))
when ((ZRange.normalize ry <=? -ZRange.normalize rny) && (ZRange.normalize rc <=? -ZRange.normalize rnc))%zrange)
; make_rewriteol
- (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ @ (-'??') @ ??))
- (fun rvc s cc rny ry y x
+ (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ' @ (-'??') @ ??))
+ (fun rvc s rcc cc rny ry y x
=> (llet2_opp2 rvc (#ident.Z_sub_get_borrow @ s @ x @ cst ry y))
- when (cc =? 0) && (ZRange.normalize ry <=? -ZRange.normalize rny)%zrange)
+ when (cc =? 0) && (ZRange.normalize ry <=? -ZRange.normalize rny)%zrange && is_bounded_by_bool cc (ZRange.normalize rcc))
; make_rewriteol
- (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ @ (-'??') @ ??))
- (fun rvc s cc rny ry y x
- => (llet2_opp2 rvc (#ident.Z_sub_with_get_borrow @ s @ ##(-cc)%Z @ x @ cst ry y))
- when (cc <? 0) && (ZRange.normalize ry <=? -ZRange.normalize rny)%zrange)
+ (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ' @ (-'??') @ ??))
+ (fun rvc s rcc cc rny ry y x
+ => (llet2_opp2 rvc (#ident.Z_sub_with_get_borrow @ s @ cst (ZRange.opp rcc) ##(-cc)%Z @ x @ cst ry y))
+ when (cc <? 0) && (ZRange.normalize ry <=? -ZRange.normalize rny)%zrange && is_bounded_by_bool cc (ZRange.normalize rcc))
; make_rewriteol
- (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ @ ?? @ (-'??')))
- (fun rvc s cc x rny ry y
+ (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ' @ ?? @ (-'??')))
+ (fun rvc s rcc cc x rny ry y
=> (llet2_opp2 rvc (#ident.Z_sub_get_borrow @ s @ x @ cst ry y))
- when (cc =? 0) && (ZRange.normalize ry <=? -ZRange.normalize rny)%zrange)
+ when (cc =? 0) && (ZRange.normalize ry <=? -ZRange.normalize rny)%zrange && is_bounded_by_bool cc (ZRange.normalize rcc))
; make_rewriteol
- (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ @ ?? @ (-'??')))
- (fun rvc s cc x rny ry y
- => (llet2_opp2 rvc (#ident.Z_sub_with_get_borrow @ s @ ##(-cc)%Z @ x @ cst ry y))
- when (cc <? 0) && (ZRange.normalize ry <=? -ZRange.normalize rny)%zrange)
+ (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ' @ ?? @ (-'??')))
+ (fun rvc s rcc cc x rny ry y
+ => (llet2_opp2 rvc (#ident.Z_sub_with_get_borrow @ s @ cst (ZRange.opp rcc) ##(-cc)%Z @ x @ cst ry y)
+ when (cc <? 0) && (ZRange.normalize ry <=? -ZRange.normalize rny)%zrange && is_bounded_by_bool cc (ZRange.normalize rcc)))
; make_rewriteol
- (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ (-'??') @ #?ℤ @ ??))
- (fun rvc s rnc rc c yy x
- => (llet2_opp2 rvc (#ident.Z_sub_with_get_borrow @ s @ cst rc c @ x @ ##(-yy)%Z))
- when (yy <=? 0) && (ZRange.normalize rc <=? -ZRange.normalize rnc)%zrange)
+ (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ (-'??') @ #?ℤ' @ ??))
+ (fun rvc s rnc rc c ryy yy x
+ => (llet2_opp2 rvc (#ident.Z_sub_with_get_borrow @ s @ cst rc c @ x @ cst (ZRange.opp ryy) ##(-yy)%Z))
+ when (yy <=? 0) && (ZRange.normalize rc <=? -ZRange.normalize rnc)%zrange && is_bounded_by_bool yy (ZRange.normalize ryy))
; make_rewriteol
- (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ (-'??') @ ?? @ #?ℤ))
- (fun rvc s rnc rc c x yy
- => (llet2_opp2 rvc (#ident.Z_sub_with_get_borrow @ s @ cst rc c @ x @ ##(-yy)%Z))
- when (yy <=? 0) && (ZRange.normalize rc <=? -ZRange.normalize rnc)%zrange)
+ (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ (-'??') @ ?? @ #?ℤ'))
+ (fun rvc s rnc rc c x ryy yy
+ => (llet2_opp2 rvc (#ident.Z_sub_with_get_borrow @ s @ cst rc c @ x @ cst (ZRange.opp ryy) ##(-yy)%Z))
+ when (yy <=? 0) && (ZRange.normalize rc <=? -ZRange.normalize rnc)%zrange && is_bounded_by_bool yy (ZRange.normalize ryy))
; make_rewriteol
- (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ @ #?ℤ @ ??))
- (fun rvc s cc yy x
- => (llet2_opp2 rvc (#ident.Z_sub_with_get_borrow @ s @ ##(-cc)%Z @ x @ ##(-yy)%Z))
- when (yy <=? 0) && (cc <=? 0) && ((yy + cc) <? 0)) (* at least one must be strictly negative *)
+ (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ' @ #?ℤ' @ ??))
+ (fun rvc s rcc cc ryy yy x
+ => (llet2_opp2 rvc (#ident.Z_sub_with_get_borrow @ s @ cst (ZRange.opp rcc) ##(-cc)%Z @ x @ cst (ZRange.opp ryy) ##(-yy)%Z))
+ when (yy <=? 0) && (cc <=? 0) && ((yy + cc) <? 0) && is_bounded_by_bool yy (ZRange.normalize ryy) && is_bounded_by_bool cc (ZRange.normalize rcc)) (* at least one must be strictly negative *)
; make_rewriteol
- (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ @ ?? @ #?ℤ))
- (fun rvc s cc x yy
- => (llet2_opp2 rvc (#ident.Z_sub_with_get_borrow @ s @ ##(-cc)%Z @ x @ ##(-yy)%Z))
- when (yy <=? 0) && (cc <=? 0) && ((yy + cc) <? 0)) (* at least one must be strictly negative *)
+ (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ' @ ?? @ #?ℤ'))
+ (fun rvc s rcc cc x ryy yy
+ => (llet2_opp2 rvc (#ident.Z_sub_with_get_borrow @ s @ cst (ZRange.opp rcc) ##(-cc)%Z @ x @ cst (ZRange.opp ryy) ##(-yy)%Z))
+ when (yy <=? 0) && (cc <=? 0) && ((yy + cc) <? 0) && is_bounded_by_bool yy (ZRange.normalize ryy) && is_bounded_by_bool cc (ZRange.normalize rcc)) (* at least one must be strictly negative *)
- ; make_rewrite
- (#pattern.ident.Z_add_get_carry @ #?ℤ @ #?ℤ @ #?ℤ)
- (fun s xx yy => ##(Z.add_get_carry_full s xx yy))
; make_rewriteo
- (#pattern.ident.Z_add_get_carry @ #?ℤ @ #?ℤ @ ??')
- (fun s xx ry y => (cst ry y, ##0) when (xx =? 0) && (ZRange.normalize ry <=? r[0~>s-1])%zrange)
+ (#pattern.ident.Z_add_get_carry @ #?ℤ' @ #?ℤ' @ #?ℤ')
+ (fun rs s rxx xx ryy yy => ##(Z.add_get_carry_full s xx yy) when is_bounded_by_bool s (ZRange.normalize rs) && is_bounded_by_bool xx (ZRange.normalize rxx) && is_bounded_by_bool yy (ZRange.normalize ryy))
+ ; make_rewriteo
+ (#pattern.ident.Z_add_get_carry @ #?ℤ' @ #?ℤ' @ ??')
+ (fun rs s rxx xx ry y => (cst ry y, cst r[0~>0] ##0) when (xx =? 0) && (ZRange.normalize ry <=? r[0~>s-1])%zrange && is_bounded_by_bool xx (ZRange.normalize rxx) && is_bounded_by_bool s (ZRange.normalize rs))
; make_rewriteo
- (#pattern.ident.Z_add_get_carry @ #?ℤ @ ??' @ #?ℤ)
- (fun s ry y xx => (cst ry y, ##0) when (xx =? 0) && (ZRange.normalize ry <=? r[0~>s-1])%zrange)
+ (#pattern.ident.Z_add_get_carry @ #?ℤ' @ ??' @ #?ℤ')
+ (fun rs s ry y rxx xx => (cst ry y, cst r[0~>0] ##0) when (xx =? 0) && (ZRange.normalize ry <=? r[0~>s-1])%zrange && is_bounded_by_bool xx (ZRange.normalize rxx) && is_bounded_by_bool s (ZRange.normalize rs))
- ; make_rewriteo (#pattern.ident.Z_add_with_carry @ #?ℤ @ ?? @ ??) (fun cc x y => x + y when cc =? 0)
+ ; make_rewriteo (#pattern.ident.Z_add_with_carry @ #?ℤ' @ ?? @ ??) (fun rcc cc x y => x + y when (cc =? 0) && is_bounded_by_bool cc (ZRange.normalize rcc))
(*; make_rewrite_step (#pattern.ident.Z_add_with_carry @ ?? @ ?? @ ??) (fun x y z => $x + $y + $z)*)
- ; make_rewrite
- (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ #?ℤ @ #?ℤ)
- (fun s cc xx yy => ##(Z.add_with_get_carry_full s cc xx yy))
; make_rewriteo
- (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ #?ℤ @ ??')
- (fun s cc xx ry y => (cst ry y, ##0) when (cc =? 0) && (xx =? 0) && (ZRange.normalize ry <=? r[0~>s-1])%zrange)
+ (#pattern.ident.Z_add_with_get_carry @ #?ℤ' @ #?ℤ' @ #?ℤ' @ #?ℤ')
+ (fun rs s rcc cc rxx xx ryy yy => ##(Z.add_with_get_carry_full s cc xx yy) when is_bounded_by_bool s (ZRange.normalize rs) && is_bounded_by_bool cc (ZRange.normalize rcc) && is_bounded_by_bool xx (ZRange.normalize rxx) && is_bounded_by_bool yy (ZRange.normalize ryy))
; make_rewriteo
- (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ ??' @ #?ℤ)
- (fun s cc ry y xx => (cst ry y, ##0) when (cc =? 0) && (xx =? 0) && (ZRange.normalize ry <=? r[0~>s-1])%zrange)
+ (#pattern.ident.Z_add_with_get_carry @ #?ℤ' @ #?ℤ' @ #?ℤ' @ ??')
+ (fun rs s rcc cc rxx xx ry y => (cst ry y, cst r[0~>0] ##0) when (cc =? 0) && (xx =? 0) && (ZRange.normalize ry <=? r[0~>s-1])%zrange && is_bounded_by_bool s (ZRange.normalize rs) && is_bounded_by_bool cc (ZRange.normalize rcc) && is_bounded_by_bool xx (ZRange.normalize rxx))
+ ; make_rewriteo
+ (#pattern.ident.Z_add_with_get_carry @ #?ℤ' @ #?ℤ' @ ??' @ #?ℤ')
+ (fun rs s rcc cc ry y rxx xx => (cst ry y, cst r[0~>0] ##0) when (cc =? 0) && (xx =? 0) && (ZRange.normalize ry <=? r[0~>s-1])%zrange && is_bounded_by_bool s (ZRange.normalize rs) && is_bounded_by_bool cc (ZRange.normalize rcc) && is_bounded_by_bool xx (ZRange.normalize rxx))
(*; make_rewriteo
(#pattern.ident.Z_add_with_get_carry @ ?? @ ?? @ #?ℤ @ #?ℤ) (fun s c xx yy => (c, ##0) when (xx =? 0) && (yy =? 0))*)
; make_rewriteol (* carry = 0: ADC x y -> ADD x y *)
- (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ @ ?? @ ??))
- (fun rvc s cc x y
+ (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ ?? @ #?ℤ' @ ?? @ ??))
+ (fun rvc s rcc cc x y
=> (llet2 rvc (#ident.Z_add_get_carry @ s @ x @ y))
- when cc =? 0)
+ when (cc =? 0) && is_bounded_by_bool cc (ZRange.normalize rcc))
; make_rewriteol (* ADC 0 0 -> (ADX 0 0, 0) *) (* except we don't do ADX, because C stringification doesn't handle it *)
- (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??' @ #?ℤ @ #?ℤ))
- (fun rvc s rc c xx yy
- => (llet vc := cst2 rvc (#ident.Z_add_with_get_carry @ ##s @ cst rc c @ ##xx @ ##yy) in
- (cst (fst rvc) (#ident.fst @ cst2 rvc ($vc)), ##0))
- when (xx =? 0) && (yy =? 0) && (ZRange.normalize rc <=? r[0~>s-1])%zrange && is_bounded_by_bool 0 (snd rvc))
+ (#pattern.ident.Z_cast2 @ (#pattern.ident.Z_add_with_get_carry @ #?ℤ' @ ??' @ #?ℤ' @ #?ℤ'))
+ (fun rvc rs s rc c rxx xx ryy yy
+ => (llet vc := cst2 rvc (#ident.Z_add_with_get_carry @ cst rs ##s @ cst rc c @ cst rxx ##xx @ cst ryy ##yy) in
+ (cst (fst rvc) (#ident.fst @ cst2 rvc ($vc)), cst r[0~>0] ##0))
+ when (xx =? 0) && (yy =? 0) && (ZRange.normalize rc <=? r[0~>s-1])%zrange && is_bounded_by_bool 0 (snd rvc) && is_bounded_by_bool s (ZRange.normalize rs) && is_bounded_by_bool xx (ZRange.normalize rxx) && is_bounded_by_bool yy (ZRange.normalize ryy))
(* let-bind any adc/sbb/mulx *)
@@ -1916,8 +1917,6 @@ Module Compilers.
(#pattern.ident.Z_cast2 @ (#pattern.ident.Z_mul_split @ ?? @ ?? @ ??))
(fun rvc s x y => llet2 rvc (#ident.Z_mul_split @ s @ x @ y))
- ; make_rewriteo (#pattern.ident.Z_cast @ #?ℤ) (fun r v => ##v when is_bounded_by_bool v r)
-
; make_rewrite_step (#pattern.ident.Z_cast2 @ (??, ??)) (fun r v1 v2 => (#(ident.Z_cast (fst r)) @ $v1, #(ident.Z_cast (snd r)) @ $v2))
; make_rewriteo
@@ -1925,22 +1924,31 @@ Module Compilers.
(fun r1 r2 x => #(ident.Z_cast r2) @ x when ZRange.is_tighter_than_bool (ZRange.normalize r2) (ZRange.normalize r1))
].
+ Definition strip_literal_casts_rewrite_rules : rewrite_rulesT
+ := [make_rewriteo (#?ℤ') (fun rx x => ##x when is_bounded_by_bool x (ZRange.normalize rx))].
+
+
Definition nbe_dtree'
:= Eval compute in @compile_rewrites ident var pattern.ident (@pattern.ident.arg_types) pattern.Raw.ident (@pattern.ident.strip_types) pattern.Raw.ident.ident_beq 100 nbe_rewrite_rules.
Definition arith_dtree'
:= Eval compute in @compile_rewrites ident var pattern.ident (@pattern.ident.arg_types) pattern.Raw.ident (@pattern.ident.strip_types) pattern.Raw.ident.ident_beq 100 (arith_rewrite_rules 0%Z (* dummy val *)).
Definition arith_with_casts_dtree'
:= Eval compute in @compile_rewrites ident var pattern.ident (@pattern.ident.arg_types) pattern.Raw.ident (@pattern.ident.strip_types) pattern.Raw.ident.ident_beq 100 arith_with_casts_rewrite_rules.
+ Definition strip_literal_casts_dtree'
+ := Eval compute in @compile_rewrites ident var pattern.ident (@pattern.ident.arg_types) pattern.Raw.ident (@pattern.ident.strip_types) pattern.Raw.ident.ident_beq 100 strip_literal_casts_rewrite_rules.
Definition nbe_dtree : decision_tree
:= Eval compute in invert_Some nbe_dtree'.
Definition arith_dtree : decision_tree
:= Eval compute in invert_Some arith_dtree'.
Definition arith_with_casts_dtree : decision_tree
:= Eval compute in invert_Some arith_with_casts_dtree'.
+ Definition strip_literal_casts_dtree : decision_tree
+ := Eval compute in invert_Some strip_literal_casts_dtree'.
Definition nbe_default_fuel := Eval compute in List.length nbe_rewrite_rules.
Definition arith_default_fuel := Eval compute in List.length (arith_rewrite_rules 0%Z (* dummy val *)).
Definition arith_with_casts_default_fuel := Eval compute in List.length arith_with_casts_rewrite_rules.
+ Definition strip_literal_casts_default_fuel := Eval compute in List.length strip_literal_casts_rewrite_rules.
Import PrimitiveHList.
(* N.B. The [combine_hlist] call MUST eta-expand
@@ -1968,12 +1976,18 @@ Module Compilers.
Definition arith_with_casts_pr1_rewrite_rules := Eval hnf in projT1 arith_with_casts_split_rewrite_rules.
Definition arith_with_casts_pr2_rewrite_rules := Eval hnf in projT2 arith_with_casts_split_rewrite_rules.
Definition arith_with_casts_all_rewrite_rules := combine_hlist (P:=rewrite_ruleTP) (arith_with_casts_pr1_rewrite_rules) arith_with_casts_pr2_rewrite_rules.
+ Definition strip_literal_casts_split_rewrite_rules := Eval cbv [split_list projT1 projT2 strip_literal_casts_rewrite_rules] in split_list strip_literal_casts_rewrite_rules.
+ Definition strip_literal_casts_pr1_rewrite_rules := Eval hnf in projT1 strip_literal_casts_split_rewrite_rules.
+ Definition strip_literal_casts_pr2_rewrite_rules := Eval hnf in projT2 strip_literal_casts_split_rewrite_rules.
+ Definition strip_literal_casts_all_rewrite_rules := combine_hlist (P:=rewrite_ruleTP) (strip_literal_casts_pr1_rewrite_rules) strip_literal_casts_pr2_rewrite_rules.
Definition nbe_rewrite_head0 do_again {t} (idc : ident t) : value_with_lets t
:= @assemble_identifier_rewriters nbe_dtree nbe_all_rewrite_rules do_again t idc.
Definition arith_rewrite_head0 max_const_val do_again {t} (idc : ident t) : value_with_lets t
:= @assemble_identifier_rewriters arith_dtree (arith_all_rewrite_rules max_const_val) do_again t idc.
Definition arith_with_casts_rewrite_head0 do_again {t} (idc : ident t) : value_with_lets t
:= @assemble_identifier_rewriters arith_with_casts_dtree arith_with_casts_all_rewrite_rules do_again t idc.
+ Definition strip_literal_casts_rewrite_head0 do_again {t} (idc : ident t) : value_with_lets t
+ := @assemble_identifier_rewriters strip_literal_casts_dtree strip_literal_casts_all_rewrite_rules do_again t idc.
Section fancy.
Context (invert_low invert_high : Z (*log2wordmax*) -> Z -> option Z)
@@ -1996,6 +2010,7 @@ Module Compilers.
Local Notation land_good := (bounds2_good ZRange.land).
Local Notation mul_good := (bounds2_good ZRange.mul).
Local Notation cc_m_good output s := (bounds1_good (ZRange.cc_m s) output).
+ Local Notation lit_good x rx := (is_bounded_by_bool x (ZRange.normalize rx)).
Definition fancy_with_casts_rewrite_rules : rewrite_rulesT
:= [
@@ -2007,42 +2022,24 @@ Module Compilers.
(Z.add_get_carry_concrete 2^256) @@ (?x, ?y) --> (add 0) @@ (y, x)
*)
make_rewriteo
- (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ @ ??' @ (pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ))))
- (fun '((r1, r2)%core) s rx x rshiftl rland ry y mask offset => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) offset) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && land_good rland ry mask && range_in_bitwidth rshiftl s && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ @ #?ℤ @ (pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ))))
- (fun '((r1, r2)%core) s xx rshiftl rland ry y mask offset => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) offset) @ (##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && land_good rland ry mask && range_in_bitwidth rshiftl s && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
+ (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ' @ ??' @ (pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ')) @ #?ℤ'))))
+ (fun '((r1, r2)%core) rs s rx x rshiftl rland ry y rmask mask roffset offset => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) offset) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && land_good rland ry mask && range_in_bitwidth rshiftl s && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s) && lit_good s rs && lit_good mask rmask && lit_good offset roffset)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ @ (pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ)) @ ??'))
- (fun '((r1, r2)%core) s rshiftl rland ry y mask offset rx x => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) offset) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && land_good rland ry mask && range_in_bitwidth rshiftl s && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ @ (pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ)) @ #?ℤ))
- (fun '((r1, r2)%core) s rshiftl rland ry y mask offset xx => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) offset) @ (##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && land_good rland ry mask && range_in_bitwidth rshiftl s && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
+ (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ' @ (pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ')) @ #?ℤ')) @ ??'))
+ (fun '((r1, r2)%core) rs s rshiftl rland ry y rmask mask roffset offset rx x => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) offset) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && land_good rland ry mask && range_in_bitwidth rshiftl s && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s) && lit_good s rs && lit_good mask rmask && lit_good offset roffset)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ @ ??' @ (pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ))))
- (fun '((r1, r2)%core) s rx x rshiftr ry y offset => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) (-offset)) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ @ #?ℤ @ (pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ))))
- (fun '((r1, r2)%core) s xx rshiftr ry y offset => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) (-offset)) @ (##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
+ (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ' @ ??' @ (pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ'))))
+ (fun '((r1, r2)%core) rs s rx x rshiftr ry y roffset offset => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) (-offset)) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s && lit_good s rs && lit_good offset roffset)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ) @ ??'))
- (fun '((r1, r2)%core) s rshiftr ry y offset rx x => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) (-offset)) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ) @ #?ℤ))
- (fun '((r1, r2)%core) s rshiftr ry y offset xx => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) (-offset)) @ (##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
+ (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ' @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ') @ ??'))
+ (fun '((r1, r2)%core) rs s rshiftr ry y roffset offset rx x => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) (-offset)) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s && lit_good s rs && lit_good offset roffset)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ @ ??' @ ??'))
- (fun '((r1, r2)%core) s rx x ry y => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) 0) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ @ ??' @ #?ℤ))
- (fun '((r1, r2)%core) s rx x yy => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) 0) @ (cst rx x, ##yy)) when (s =? 2^Z.log2 s) && range_in_bitwidth yy s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ @ #?ℤ @ ??'))
- (fun '((r1, r2)%core) s xx ry y => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) 0) @ (##xx, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s)
+ (pcst2 (#pattern.ident.Z_add_get_carry @ #?ℤ' @ ??' @ ??'))
+ (fun '((r1, r2)%core) rs s rx x ry y => cst2 (r1, r2)%core (#(ident.fancy_add (Z.log2 s) 0) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s && lit_good s rs)
(*
(Z.add_with_get_carry_concrete 2^256) @@ (?c, ?x, ?y << 128) --> (addc 128) @@ (c, x, y)
(Z.add_with_get_carry_concrete 2^256) @@ (?c, ?x << 128, ?y) --> (addc 128) @@ (c, y, x)
@@ -2051,212 +2048,78 @@ Module Compilers.
(Z.add_with_get_carry_concrete 2^256) @@ (?c, ?x, ?y) --> (addc 0) @@ (c, y, x)
*)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??' @ ??' @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ)))
- (fun '((r1, r2)%core) s rc c rx x rshiftl rland ry y mask offset => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) offset) @ (cst rc c, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && land_good rland ry mask && range_in_bitwidth rshiftl s && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ ??' @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ)))
- (fun '((r1, r2)%core) s cc rx x rshiftl rland ry y mask offset => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) offset) @ (##cc, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??' @ #?ℤ @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ)))
- (fun '((r1, r2)%core) s rc c xx rshiftl rland ry y mask offset => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) offset) @ (cst rc c, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ #?ℤ @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ)))
- (fun '((r1, r2)%core) s cc xx rshiftl rland ry y mask offset => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) offset) @ (##cc, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
+ (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ' @ ??' @ ??' @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ')) @ #?ℤ')))
+ (fun '((r1, r2)%core) rs s rc c rx x rshiftl rland ry y rmask mask roffset offset => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) offset) @ (cst rc c, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && land_good rland ry mask && range_in_bitwidth rshiftl s && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s) && lit_good s rs && lit_good mask rmask && lit_good offset roffset)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??' @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ) @ ??'))
- (fun '((r1, r2)%core) s rc c rshiftl rland ry y mask offset rx x => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) offset) @ (cst rc c, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ) @ ??'))
- (fun '((r1, r2)%core) s cc rshiftl rland ry y mask offset rx x => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) offset) @ (##cc, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??' @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ) @ #?ℤ))
- (fun '((r1, r2)%core) s rc c rshiftl rland ry y mask offset xx => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) offset) @ (cst rc c, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ) @ #?ℤ))
- (fun '((r1, r2)%core) s cc rshiftl rland ry y mask offset xx => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) offset) @ (##cc, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
+ (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ' @ ??' @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ')) @ #?ℤ') @ ??'))
+ (fun '((r1, r2)%core) rs s rc c rshiftl rland ry y rmask mask roffset offset rx x => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) offset) @ (cst rc c, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s) && lit_good s rs && lit_good mask rmask && lit_good offset roffset)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??' @ ??' @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ)))
- (fun '((r1, r2)%core) s rc c rx x rshiftr ry y offset => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) (-offset)) @ (cst rc c, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ ??' @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ)))
- (fun '((r1, r2)%core) s cc rx x rshiftr ry y offset => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) (-offset)) @ (##cc, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??' @ #?ℤ @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ)))
- (fun '((r1, r2)%core) s rc c xx rshiftr ry y offset => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) (-offset)) @ (cst rc c, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ #?ℤ @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ)))
- (fun '((r1, r2)%core) s cc xx rshiftr ry y offset => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) (-offset)) @ (##cc, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
+ (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ' @ ??' @ ??' @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ')))
+ (fun '((r1, r2)%core) rs s rc c rx x rshiftr ry y roffset offset => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) (-offset)) @ (cst rc c, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s && lit_good s rs && lit_good offset roffset)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??' @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ) @ ??'))
- (fun '((r1, r2)%core) s rc c rshiftr ry y offset rx x => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) (-offset)) @ (cst rc c, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ) @ ??'))
- (fun '((r1, r2)%core) s cc rshiftr ry y offset rx x => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) (-offset)) @ (##cc, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??' @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ) @ #?ℤ))
- (fun '((r1, r2)%core) s rc c rshiftr ry y offset xx => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) (-offset)) @ (cst rc c, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ) @ #?ℤ))
- (fun '((r1, r2)%core) s cc rshiftr ry y offset xx => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) (-offset)) @ (##cc, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
+ (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ' @ ??' @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ') @ ??'))
+ (fun '((r1, r2)%core) rs s rc c rshiftr ry y roffset offset rx x => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) (-offset)) @ (cst rc c, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s && lit_good s rs && lit_good offset roffset)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??' @ ??' @ ??'))
- (fun '((r1, r2)%core) s rc c rx x ry y => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) 0) @ (cst rc c, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ ??' @ ??'))
- (fun '((r1, r2)%core) s cc rx x ry y => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) 0) @ (##cc, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??' @ #?ℤ @ ??'))
- (fun '((r1, r2)%core) s rc c xx ry y => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) 0) @ (cst rc c, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??' @ ??' @ #?ℤ))
- (fun '((r1, r2)%core) s rc c rx x yy => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) 0) @ (cst rc c, cst rx x, ##yy)) when (s =? 2^Z.log2 s) && range_in_bitwidth yy s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ #?ℤ @ ??'))
- (fun '((r1, r2)%core) s cc xx ry y => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) 0) @ (##cc, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ ??' @ #?ℤ))
- (fun '((r1, r2)%core) s cc rx x yy => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) 0) @ (##cc, cst rx x, ##yy)) when (s =? 2^Z.log2 s) && range_in_bitwidth yy s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ @ ??' @ #?ℤ @ #?ℤ))
- (fun '((r1, r2)%core) s rc c xx yy => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) 0) @ (cst rc c, ##xx, ##yy)) when (s =? 2^Z.log2 s) && range_in_bitwidth yy s)
+ (pcst2 (#pattern.ident.Z_add_with_get_carry @ #?ℤ' @ ??' @ ??' @ ??'))
+ (fun '((r1, r2)%core) rs s rc c rx x ry y => cst2 (r1, r2)%core (#(ident.fancy_addc (Z.log2 s) 0) @ (cst rc c, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s && lit_good s rs)
+
(*
(Z.sub_get_borrow_concrete 2^256) @@ (?x, ?y << 128) --> (sub 128) @@ (x, y)
(Z.sub_get_borrow_concrete 2^256) @@ (?x, ?y >> 128) --> (sub (- 128)) @@ (x, y)
(Z.sub_get_borrow_concrete 2^256) @@ (?x, ?y) --> (sub 0) @@ (y, x)
*)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_get_borrow @ #?ℤ @ ??' @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ)))
- (fun '((r1, r2)%core) s rx x rshiftl rland ry y mask offset => cst2 (r1, r2)%core (#(ident.fancy_sub (Z.log2 s) offset) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_get_borrow @ #?ℤ @ #?ℤ @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ)))
- (fun '((r1, r2)%core) s xx rshiftl rland ry y mask offset => cst2 (r1, r2)%core (#(ident.fancy_sub (Z.log2 s) offset) @ (##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
+ (pcst2 (#pattern.ident.Z_sub_get_borrow @ #?ℤ' @ ??' @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ')) @ #?ℤ')))
+ (fun '((r1, r2)%core) rs s rx x rshiftl rland ry y rmask mask roffset offset => cst2 (r1, r2)%core (#(ident.fancy_sub (Z.log2 s) offset) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s) && lit_good s rs && lit_good mask rmask && lit_good offset roffset)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_get_borrow @ #?ℤ @ ??' @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ)))
- (fun '((r1, r2)%core) s rx x rshiftr ry y offset => cst2 (r1, r2)%core (#(ident.fancy_sub (Z.log2 s) (-offset)) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_get_borrow @ #?ℤ @ #?ℤ @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ)))
- (fun '((r1, r2)%core) s xx rshiftr ry y offset => cst2 (r1, r2)%core (#(ident.fancy_sub (Z.log2 s) (-offset)) @ (##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
+ (pcst2 (#pattern.ident.Z_sub_get_borrow @ #?ℤ' @ ??' @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ')))
+ (fun '((r1, r2)%core) rs s rx x rshiftr ry y roffset offset => cst2 (r1, r2)%core (#(ident.fancy_sub (Z.log2 s) (-offset)) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s && lit_good s rs && lit_good offset roffset)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_get_borrow @ #?ℤ @ ??' @ ??'))
- (fun '((r1, r2)%core) s rx x ry y => cst2 (r1, r2)%core (#(ident.fancy_sub (Z.log2 s) 0) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_get_borrow @ #?ℤ @ #?ℤ @ ??'))
- (fun '((r1, r2)%core) s xx ry y => cst2 (r1, r2)%core (#(ident.fancy_sub (Z.log2 s) 0) @ (##xx, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_get_borrow @ #?ℤ @ ??' @ #?ℤ))
- (fun '((r1, r2)%core) s rx x yy => cst2 (r1, r2)%core (#(ident.fancy_sub (Z.log2 s) 0) @ (cst rx x, ##yy)) when (s =? 2^Z.log2 s) && range_in_bitwidth yy s)
+ (pcst2 (#pattern.ident.Z_sub_get_borrow @ #?ℤ' @ ??' @ ??'))
+ (fun '((r1, r2)%core) rs s rx x ry y => cst2 (r1, r2)%core (#(ident.fancy_sub (Z.log2 s) 0) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s && lit_good s rs)
(*
(Z.sub_with_get_borrow_concrete 2^256) @@ (?c, ?x, ?y << 128) --> (subb 128) @@ (c, x, y)
(Z.sub_with_get_borrow_concrete 2^256) @@ (?c, ?x, ?y >> 128) --> (subb (- 128)) @@ (c, x, y)
(Z.sub_with_get_borrow_concrete 2^256) @@ (?c, ?x, ?y) --> (subb 0) @@ (c, y, x)
*)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ ??' @ ??' @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ)))
- (fun '((r1, r2)%core) s rb b rx x rshiftl rland ry y mask offset => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) offset) @ (cst rb b, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ #?ℤ @ ??' @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ)))
- (fun '((r1, r2)%core) s bb rx x rshiftl rland ry y mask offset => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) offset) @ (##bb, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ ??' @ #?ℤ @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ)))
- (fun '((r1, r2)%core) s rb b xx rshiftl rland ry y mask offset => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) offset) @ (cst rb b, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ #?ℤ @ #?ℤ @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ)) @ #?ℤ)))
- (fun '((r1, r2)%core) s bb xx rshiftl rland ry y mask offset => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) offset) @ (##bb, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s))
+ (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ' @ ??' @ ??' @ pcst (#pattern.ident.Z_shiftl @ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ')) @ #?ℤ')))
+ (fun '((r1, r2)%core) rs s rb b rx x rshiftl rland ry y rmask mask roffset offset => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) offset) @ (cst rb b, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftl_good rshiftl rland offset && range_in_bitwidth rshiftl s && land_good rland ry mask && (mask =? Z.ones (Z.log2 s - offset)) && (0 <=? offset) && (offset <=? Z.log2 s) && lit_good s rs && lit_good mask rmask && lit_good offset roffset)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ ??' @ ??' @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ)))
- (fun '((r1, r2)%core) s rb b rx x rshiftr ry y offset => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) (-offset)) @ (cst rb b, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ #?ℤ @ ??' @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ)))
- (fun '((r1, r2)%core) s bb rx x rshiftr ry y offset => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) (-offset)) @ (##bb, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ ??' @ #?ℤ @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ)))
- (fun '((r1, r2)%core) s rb b xx rshiftr ry y offset => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) (-offset)) @ (cst rb b, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ #?ℤ @ #?ℤ @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ)))
- (fun '((r1, r2)%core) s bb xx rshiftr ry y offset => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) (-offset)) @ (##bb, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s)
+ (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ' @ ??' @ ??' @ pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ')))
+ (fun '((r1, r2)%core) rs s rb b rx x rshiftr ry y roffset offset => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) (-offset)) @ (cst rb b, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && shiftr_good rshiftr ry offset && range_in_bitwidth rshiftr s && lit_good s rs && lit_good offset roffset)
; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ ??' @ ??' @ ??'))
- (fun '((r1, r2)%core) s rb b rx x ry y => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) 0) @ (cst rb b, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ #?ℤ @ ??' @ ??'))
- (fun '((r1, r2)%core) s bb rx x ry y => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) 0) @ (##bb, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ ??' @ #?ℤ @ ??'))
- (fun '((r1, r2)%core) s rb b xx ry y => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) 0) @ (cst rb b, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ ??' @ ??' @ #?ℤ))
- (fun '((r1, r2)%core) s rb b rx x yy => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) 0) @ (cst rb b, cst rx x, ##yy)) when (s =? 2^Z.log2 s) && range_in_bitwidth yy s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ #?ℤ @ #?ℤ @ ??'))
- (fun '((r1, r2)%core) s bb xx ry y => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) 0) @ (##bb, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ #?ℤ @ ??' @ #?ℤ))
- (fun '((r1, r2)%core) s bb rx x yy => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) 0) @ (##bb, cst rx x, ##yy)) when (s =? 2^Z.log2 s) && range_in_bitwidth yy s)
- ; make_rewriteo
- (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ @ ??' @ #?ℤ @ #?ℤ))
- (fun '((r1, r2)%core) s rb b xx yy => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) 0) @ (cst rb b, ##xx, ##yy)) when (s =? 2^Z.log2 s) && range_in_bitwidth yy s)
+ (pcst2 (#pattern.ident.Z_sub_with_get_borrow @ #?ℤ' @ ??' @ ??' @ ??'))
+ (fun '((r1, r2)%core) rs s rb b rx x ry y => cst2 (r1, r2)%core (#(ident.fancy_subb (Z.log2 s) 0) @ (cst rb b, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && range_in_bitwidth ry s && lit_good s rs)
(*(Z.rshi_concrete 2^256 ?n) @@ (?c, ?x, ?y) --> (rshi n) @@ (x, y)*)
; make_rewriteo
- (pcst (#pattern.ident.Z_rshi @ #?ℤ @ ??' @ ??' @ #?ℤ))
- (fun r s rx x ry y n => cst r (#(ident.fancy_rshi (Z.log2 s) n) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s))
- ; make_rewriteo
- (pcst (#pattern.ident.Z_rshi @ #?ℤ @ #?ℤ @ ??' @ #?ℤ))
- (fun r s xx ry y n => cst r (#(ident.fancy_rshi (Z.log2 s) n) @ (##xx, cst ry y)) when (s =? 2^Z.log2 s))
- ; make_rewriteo
- (pcst (#pattern.ident.Z_rshi @ #?ℤ @ ??' @ #?ℤ @ #?ℤ))
- (fun r s rx x yy n => cst r (#(ident.fancy_rshi (Z.log2 s) n) @ (cst rx x, ##yy)) when (s =? 2^Z.log2 s))
+ (pcst (#pattern.ident.Z_rshi @ #?ℤ' @ ??' @ ??' @ #?ℤ'))
+ (fun r rs s rx x ry y rn n => cst r (#(ident.fancy_rshi (Z.log2 s) n) @ (cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && lit_good s rs && lit_good n rn)
(*
Z.zselect @@ (Z.cc_m_concrete 2^256 ?c, ?x, ?y) --> selm @@ (c, x, y)
Z.zselect @@ (?c &' 1, ?x, ?y) --> sell @@ (c, x, y)
Z.zselect @@ (?c, ?x, ?y) --> selc @@ (c, x, y)
*)
; make_rewriteo
- (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_cc_m @ #?ℤ @ ??') @ ??' @ ??'))
- (fun r rccm s rc c rx x ry y => cst r (#(ident.fancy_selm (Z.log2 s)) @ (cst rc c, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && cc_m_good rccm s rc)
- ; make_rewriteo
- (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_cc_m @ #?ℤ @ ??') @ #?ℤ @ ??'))
- (fun r rccm s rc c xx ry y => cst r (#(ident.fancy_selm (Z.log2 s)) @ (cst rc c, ##xx, cst ry y)) when (s =? 2^Z.log2 s) && cc_m_good rccm s rc)
- ; make_rewriteo
- (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_cc_m @ #?ℤ @ ??') @ ??' @ #?ℤ))
- (fun r rccm s rc c rx x yy => cst r (#(ident.fancy_selm (Z.log2 s)) @ (cst rc c, cst rx x, ##yy)) when (s =? 2^Z.log2 s) && cc_m_good rccm s rc)
- ; make_rewriteo
- (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_cc_m @ #?ℤ @ ??') @ #?ℤ @ #?ℤ))
- (fun r rccm s rc c xx yy => cst r (#(ident.fancy_selm (Z.log2 s)) @ (cst rc c, ##xx, ##yy)) when (s =? 2^Z.log2 s) && cc_m_good rccm s rc)
+ (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_cc_m @ #?ℤ' @ ??') @ ??' @ ??'))
+ (fun r rccm rs s rc c rx x ry y => cst r (#(ident.fancy_selm (Z.log2 s)) @ (cst rc c, cst rx x, cst ry y)) when (s =? 2^Z.log2 s) && cc_m_good rccm s rc && lit_good s rs)
; make_rewriteo
- (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_land @ #?ℤ @ ??') @ ??' @ ??'))
- (fun r rland mask rc c rx x ry y => cst r (#ident.fancy_sell @ (cst rc c, cst rx x, cst ry y)) when (mask =? 1) && land_good rland mask rc)
- ; make_rewriteo
- (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_land @ #?ℤ @ ??') @ #?ℤ @ ??'))
- (fun r rland mask rc c xx ry y => cst r (#ident.fancy_sell @ (cst rc c, ##xx, cst ry y)) when (mask =? 1) && land_good rland mask rc)
- ; make_rewriteo
- (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_land @ #?ℤ @ ??') @ ??' @ #?ℤ))
- (fun r rland mask rc c rx x yy => cst r (#ident.fancy_sell @ (cst rc c, cst rx x, ##yy)) when (mask =? 1) && land_good rland mask rc)
- ; make_rewriteo
- (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_land @ #?ℤ @ ??') @ #?ℤ @ #?ℤ))
- (fun r rland mask rc c xx yy => cst r (#ident.fancy_sell @ (cst rc c, ##xx, ##yy)) when (mask =? 1) && land_good rland mask rc)
+ (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_land @ #?ℤ' @ ??') @ ??' @ ??'))
+ (fun r rland rmask mask rc c rx x ry y => cst r (#ident.fancy_sell @ (cst rc c, cst rx x, cst ry y)) when (mask =? 1) && land_good rland mask rc && lit_good mask rmask)
; make_rewriteo
- (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_land @ ??' @ #?ℤ) @ ??' @ ??'))
- (fun r rland rc c mask rx x ry y => cst r (#ident.fancy_sell @ (cst rc c, cst rx x, cst ry y)) when (mask =? 1) && land_good rland rc mask)
- ; make_rewriteo
- (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_land @ ??' @ #?ℤ) @ #?ℤ @ ??'))
- (fun r rland rc c mask xx ry y => cst r (#ident.fancy_sell @ (cst rc c, ##xx, cst ry y)) when (mask =? 1) && land_good rland rc mask)
- ; make_rewriteo
- (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_land @ ??' @ #?ℤ) @ ??' @ #?ℤ))
- (fun r rland rc c mask rx x yy => cst r (#ident.fancy_sell @ (cst rc c, cst rx x, ##yy)) when (mask =? 1) && land_good rland rc mask)
- ; make_rewriteo
- (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_land @ ??' @ #?ℤ) @ #?ℤ @ #?ℤ))
- (fun r rland rc c mask xx yy => cst r (#ident.fancy_sell @ (cst rc c, ##xx, ##yy)) when (mask =? 1) && land_good rland rc mask)
+ (pcst (#pattern.ident.Z_zselect @ pcst (#pattern.ident.Z_land @ ??' @ #?ℤ') @ ??' @ ??'))
+ (fun r rland rc c rmask mask rx x ry y => cst r (#ident.fancy_sell @ (cst rc c, cst rx x, cst ry y)) when (mask =? 1) && land_good rland rc mask && lit_good mask rmask)
; make_rewrite
(pcst (#pattern.ident.Z_zselect @ ?? @ ?? @ ??))
@@ -2274,70 +2137,70 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
*)
(* literal on left *)
; make_rewriteo
- (#?ℤ *' pcst (#pattern.ident.Z_land @ ??' @ #?ℤ))
- (fun r x rland ry y mask => let s := (2*Z.log2_up mask)%Z in x <- invert_low s x; cst r (#(ident.fancy_mulll s) @ (##x, cst ry y)) when (mask =? 2^(s/2)-1) && land_good rland ry mask)
+ (#?ℤ' *' pcst (#pattern.ident.Z_land @ ??' @ #?ℤ'))
+ (fun r rx x rland ry y rmask mask => let s := (2*Z.log2_up mask)%Z in x <- invert_low s x; cst r (#(ident.fancy_mulll s) @ (##x, cst ry y)) when (mask =? 2^(s/2)-1) && land_good rland ry mask && lit_good x rx && lit_good mask rmask)
; make_rewriteo
- (#?ℤ *' pcst (#pattern.ident.Z_land @ #?ℤ @ ??'))
- (fun r x rland mask ry y => let s := (2*Z.log2_up mask)%Z in x <- invert_low s x; cst r (#(ident.fancy_mulll s) @ (##x, cst ry y)) when (mask =? 2^(s/2)-1) && land_good rland mask ry)
+ (#?ℤ' *' pcst (#pattern.ident.Z_land @ #?ℤ' @ ??'))
+ (fun r rx x rland rmask mask ry y => let s := (2*Z.log2_up mask)%Z in x <- invert_low s x; cst r (#(ident.fancy_mulll s) @ (##x, cst ry y)) when (mask =? 2^(s/2)-1) && land_good rland mask ry && lit_good x rx && lit_good mask rmask)
; make_rewriteo
- (#?ℤ *' pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ))
- (fun r x rshiftr ry y offset => let s := (2*offset)%Z in x <- invert_low s x; cst r (#(ident.fancy_mullh s) @ (##x, cst ry y)) when shiftr_good rshiftr ry offset)
+ (#?ℤ' *' pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ'))
+ (fun r rx x rshiftr ry y roffset offset => let s := (2*offset)%Z in x <- invert_low s x; cst r (#(ident.fancy_mullh s) @ (##x, cst ry y)) when shiftr_good rshiftr ry offset && lit_good x rx && lit_good offset roffset)
; make_rewriteo
- (#?ℤ *' pcst (#pattern.ident.Z_land @ #?ℤ @ ??'))
- (fun r x rland mask ry y => let s := (2*Z.log2_up mask)%Z in x <- invert_high s x; cst r (#(ident.fancy_mulhl s) @ (##x, cst ry y)) when (mask =? 2^(s/2)-1) && land_good rland mask ry)
+ (#?ℤ' *' pcst (#pattern.ident.Z_land @ #?ℤ' @ ??'))
+ (fun r rx x rland rmask mask ry y => let s := (2*Z.log2_up mask)%Z in x <- invert_high s x; cst r (#(ident.fancy_mulhl s) @ (##x, cst ry y)) when (mask =? 2^(s/2)-1) && land_good rland mask ry && lit_good x rx && lit_good mask rmask)
; make_rewriteo
- (#?ℤ *' pcst (#pattern.ident.Z_land @ ??' @ #?ℤ))
- (fun r x rland ry y mask => let s := (2*Z.log2_up mask)%Z in x <- invert_high s x; cst r (#(ident.fancy_mulhl s) @ (##x, cst ry y)) when (mask =? 2^(s/2)-1) && land_good rland ry mask)
+ (#?ℤ' *' pcst (#pattern.ident.Z_land @ ??' @ #?ℤ'))
+ (fun r rx x rland ry y rmask mask => let s := (2*Z.log2_up mask)%Z in x <- invert_high s x; cst r (#(ident.fancy_mulhl s) @ (##x, cst ry y)) when (mask =? 2^(s/2)-1) && land_good rland ry mask && lit_good x rx && lit_good mask rmask)
; make_rewriteo
- (#?ℤ *' pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ))
- (fun r x rshiftr ry y offset => let s := (2*offset)%Z in x <- invert_high s x; cst r (#(ident.fancy_mulhh s) @ (##x, cst ry y)) when shiftr_good rshiftr ry offset)
+ (#?ℤ' *' pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ'))
+ (fun r rx x rshiftr ry y roffset offset => let s := (2*offset)%Z in x <- invert_high s x; cst r (#(ident.fancy_mulhh s) @ (##x, cst ry y)) when shiftr_good rshiftr ry offset && lit_good x rx && lit_good offset roffset)
(* literal on right *)
; make_rewriteo
- (pcst (#pattern.ident.Z_land @ #?ℤ @ ??') *' #?ℤ)
- (fun r rland mask rx x y => let s := (2*Z.log2_up mask)%Z in y <- invert_low s y; cst r (#(ident.fancy_mulll s) @ (cst rx x, ##y)) when (mask =? 2^(s/2)-1) && land_good rland mask rx)
+ (pcst (#pattern.ident.Z_land @ #?ℤ' @ ??') *' #?ℤ')
+ (fun r rland rmask mask rx x ry y => let s := (2*Z.log2_up mask)%Z in y <- invert_low s y; cst r (#(ident.fancy_mulll s) @ (cst rx x, ##y)) when (mask =? 2^(s/2)-1) && land_good rland mask rx && lit_good y ry && lit_good mask rmask)
; make_rewriteo
- (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ) *' #?ℤ)
- (fun r rland rx x mask y => let s := (2*Z.log2_up mask)%Z in y <- invert_low s y; cst r (#(ident.fancy_mulll s) @ (cst rx x, ##y)) when (mask =? 2^(s/2)-1) && land_good rland rx mask)
+ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ') *' #?ℤ')
+ (fun r rland rx x rmask mask ry y => let s := (2*Z.log2_up mask)%Z in y <- invert_low s y; cst r (#(ident.fancy_mulll s) @ (cst rx x, ##y)) when (mask =? 2^(s/2)-1) && land_good rland rx mask && lit_good y ry && lit_good mask rmask)
; make_rewriteo
- (pcst (#pattern.ident.Z_land @ #?ℤ @ ??') *' #?ℤ)
- (fun r rland mask rx x y => let s := (2*Z.log2_up mask)%Z in y <- invert_high s y; cst r (#(ident.fancy_mullh s) @ (cst rx x, ##y)) when (mask =? 2^(s/2)-1) && land_good rland mask rx)
+ (pcst (#pattern.ident.Z_land @ #?ℤ' @ ??') *' #?ℤ')
+ (fun r rland rmask mask rx x ry y => let s := (2*Z.log2_up mask)%Z in y <- invert_high s y; cst r (#(ident.fancy_mullh s) @ (cst rx x, ##y)) when (mask =? 2^(s/2)-1) && land_good rland mask rx && lit_good y ry && lit_good mask rmask)
; make_rewriteo
- (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ) *' #?ℤ)
- (fun r rland rx x mask y => let s := (2*Z.log2_up mask)%Z in y <- invert_high s y; cst r (#(ident.fancy_mullh s) @ (cst rx x, ##y)) when (mask =? 2^(s/2)-1) && land_good rland rx mask)
+ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ') *' #?ℤ')
+ (fun r rland rx x rmask mask ry y => let s := (2*Z.log2_up mask)%Z in y <- invert_high s y; cst r (#(ident.fancy_mullh s) @ (cst rx x, ##y)) when (mask =? 2^(s/2)-1) && land_good rland rx mask && lit_good y ry && lit_good mask rmask)
; make_rewriteo
- (pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ) *' #?ℤ)
- (fun r rshiftr rx x offset y => let s := (2*offset)%Z in y <- invert_low s y; cst r (#(ident.fancy_mulhl s) @ (cst rx x, ##y)) when shiftr_good rshiftr rx offset)
+ (pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ') *' #?ℤ')
+ (fun r rshiftr rx x roffset offset ry y => let s := (2*offset)%Z in y <- invert_low s y; cst r (#(ident.fancy_mulhl s) @ (cst rx x, ##y)) when shiftr_good rshiftr rx offset && lit_good y ry && lit_good offset roffset)
; make_rewriteo
- (pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ) *' #?ℤ)
- (fun r rshiftr rx x offset y => let s := (2*offset)%Z in y <- invert_high s y; cst r (#(ident.fancy_mulhh s) @ (cst rx x, ##y)) when shiftr_good rshiftr rx offset)
+ (pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ') *' #?ℤ')
+ (fun r rshiftr rx x roffset offset ry y => let s := (2*offset)%Z in y <- invert_high s y; cst r (#(ident.fancy_mulhh s) @ (cst rx x, ##y)) when shiftr_good rshiftr rx offset && lit_good y ry && lit_good offset roffset)
(* no literal *)
; make_rewriteo
- (pcst (#pattern.ident.Z_land @ #?ℤ @ ??') *' pcst (#pattern.ident.Z_land @ #?ℤ @ ??'))
- (fun r rland1 mask1 rx x rland2 mask2 ry y => let s := (2*Z.log2_up mask1)%Z in cst r (#(ident.fancy_mulll s) @ (cst rx x, cst ry y)) when (mask1 =? 2^(s/2)-1) && (mask2 =? 2^(s/2)-1) && land_good rland1 mask1 rx && land_good rland2 mask2 ry)
+ (pcst (#pattern.ident.Z_land @ #?ℤ' @ ??') *' pcst (#pattern.ident.Z_land @ #?ℤ' @ ??'))
+ (fun r rland1 rmask1 mask1 rx x rland2 rmask2 mask2 ry y => let s := (2*Z.log2_up mask1)%Z in cst r (#(ident.fancy_mulll s) @ (cst rx x, cst ry y)) when (mask1 =? 2^(s/2)-1) && (mask2 =? 2^(s/2)-1) && land_good rland1 mask1 rx && land_good rland2 mask2 ry && lit_good mask1 rmask1 && lit_good mask2 rmask2)
; make_rewriteo
- (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ) *' pcst (#pattern.ident.Z_land @ #?ℤ @ ??'))
- (fun r rland1 rx x mask1 rland2 mask2 ry y => let s := (2*Z.log2_up mask1)%Z in cst r (#(ident.fancy_mulll s) @ (cst rx x, cst ry y)) when (mask1 =? 2^(s/2)-1) && (mask2 =? 2^(s/2)-1) && land_good rland1 rx mask1 && land_good rland2 mask2 ry)
+ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ') *' pcst (#pattern.ident.Z_land @ #?ℤ' @ ??'))
+ (fun r rland1 rx x rmask1 mask1 rland2 rmask2 mask2 ry y => let s := (2*Z.log2_up mask1)%Z in cst r (#(ident.fancy_mulll s) @ (cst rx x, cst ry y)) when (mask1 =? 2^(s/2)-1) && (mask2 =? 2^(s/2)-1) && land_good rland1 rx mask1 && land_good rland2 mask2 ry && lit_good mask1 rmask1 && lit_good mask2 rmask2)
; make_rewriteo
- (pcst (#pattern.ident.Z_land @ #?ℤ @ ??') *' pcst (#pattern.ident.Z_land @ ??' @ #?ℤ))
- (fun r rland1 mask1 rx x rland2 ry y mask2 => let s := (2*Z.log2_up mask1)%Z in cst r (#(ident.fancy_mulll s) @ (cst rx x, cst ry y)) when (mask1 =? 2^(s/2)-1) && (mask2 =? 2^(s/2)-1) && land_good rland1 mask1 rx && land_good rland2 ry mask2)
+ (pcst (#pattern.ident.Z_land @ #?ℤ' @ ??') *' pcst (#pattern.ident.Z_land @ ??' @ #?ℤ'))
+ (fun r rland1 rmask1 mask1 rx x rland2 ry y rmask2 mask2 => let s := (2*Z.log2_up mask1)%Z in cst r (#(ident.fancy_mulll s) @ (cst rx x, cst ry y)) when (mask1 =? 2^(s/2)-1) && (mask2 =? 2^(s/2)-1) && land_good rland1 mask1 rx && land_good rland2 ry mask2 && lit_good mask1 rmask1 && lit_good mask2 rmask2)
; make_rewriteo
- (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ) *' pcst (#pattern.ident.Z_land @ ??' @ #?ℤ))
- (fun r rland1 rx x mask1 rland2 ry y mask2 => let s := (2*Z.log2_up mask1)%Z in cst r (#(ident.fancy_mulll s) @ (cst rx x, cst ry y)) when (mask1 =? 2^(s/2)-1) && (mask2 =? 2^(s/2)-1) && land_good rland1 rx mask1 && land_good rland2 ry mask2)
+ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ') *' pcst (#pattern.ident.Z_land @ ??' @ #?ℤ'))
+ (fun r rland1 rx x rmask1 mask1 rland2 ry y rmask2 mask2 => let s := (2*Z.log2_up mask1)%Z in cst r (#(ident.fancy_mulll s) @ (cst rx x, cst ry y)) when (mask1 =? 2^(s/2)-1) && (mask2 =? 2^(s/2)-1) && land_good rland1 rx mask1 && land_good rland2 ry mask2 && lit_good mask1 rmask1 && lit_good mask2 rmask2)
; make_rewriteo
- (pcst (#pattern.ident.Z_land @ #?ℤ @ ??') *' pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ))
- (fun r rland1 mask rx x rshiftr2 ry y offset => let s := (2*offset)%Z in cst r (#(ident.fancy_mullh s) @ (cst rx x, cst ry y)) when (mask =? 2^(s/2)-1) && land_good rland1 mask rx && shiftr_good rshiftr2 ry offset)
+ (pcst (#pattern.ident.Z_land @ #?ℤ' @ ??') *' pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ'))
+ (fun r rland1 rmask mask rx x rshiftr2 ry y roffset offset => let s := (2*offset)%Z in cst r (#(ident.fancy_mullh s) @ (cst rx x, cst ry y)) when (mask =? 2^(s/2)-1) && land_good rland1 mask rx && shiftr_good rshiftr2 ry offset && lit_good mask rmask && lit_good offset roffset)
; make_rewriteo
- (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ) *' pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ))
- (fun r rland1 rx x mask rshiftr2 ry y offset => let s := (2*offset)%Z in cst r (#(ident.fancy_mullh s) @ (cst rx x, cst ry y)) when (mask =? 2^(s/2)-1) && land_good rland1 rx mask && shiftr_good rshiftr2 ry offset)
+ (pcst (#pattern.ident.Z_land @ ??' @ #?ℤ') *' pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ'))
+ (fun r rland1 rx x rmask mask rshiftr2 ry y roffset offset => let s := (2*offset)%Z in cst r (#(ident.fancy_mullh s) @ (cst rx x, cst ry y)) when (mask =? 2^(s/2)-1) && land_good rland1 rx mask && shiftr_good rshiftr2 ry offset && lit_good mask rmask && lit_good offset roffset)
; make_rewriteo
- (pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ) *' pcst (#pattern.ident.Z_land @ #?ℤ @ ??'))
- (fun r rshiftr1 rx x offset rland2 mask ry y => let s := (2*offset)%Z in cst r (#(ident.fancy_mulhl s) @ (cst rx x, cst ry y)) when (mask =? 2^(s/2)-1) && shiftr_good rshiftr1 rx offset && land_good rland2 mask ry)
+ (pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ') *' pcst (#pattern.ident.Z_land @ #?ℤ' @ ??'))
+ (fun r rshiftr1 rx x roffset offset rland2 rmask mask ry y => let s := (2*offset)%Z in cst r (#(ident.fancy_mulhl s) @ (cst rx x, cst ry y)) when (mask =? 2^(s/2)-1) && shiftr_good rshiftr1 rx offset && land_good rland2 mask ry && lit_good mask rmask && lit_good offset roffset)
; make_rewriteo
- (pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ) *' pcst (#pattern.ident.Z_land @ ??' @ #?ℤ))
- (fun r rshiftr1 rx x offset rland2 ry y mask => let s := (2*offset)%Z in cst r (#(ident.fancy_mulhl s) @ (cst rx x, cst ry y)) when (mask =? 2^(s/2)-1) && shiftr_good rshiftr1 rx offset && land_good rland2 ry mask)
+ (pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ') *' pcst (#pattern.ident.Z_land @ ??' @ #?ℤ'))
+ (fun r rshiftr1 rx x roffset offset rland2 ry y rmask mask => let s := (2*offset)%Z in cst r (#(ident.fancy_mulhl s) @ (cst rx x, cst ry y)) when (mask =? 2^(s/2)-1) && shiftr_good rshiftr1 rx offset && land_good rland2 ry mask && lit_good mask rmask && lit_good offset roffset)
; make_rewriteo
- (pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ) *' pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ))
- (fun r rshiftr1 rx x offset1 rshiftr2 ry y offset2 => let s := (2*offset1)%Z in cst r (#(ident.fancy_mulhh s) @ (cst rx x, cst ry y)) when (offset1 =? offset2) && shiftr_good rshiftr1 rx offset1 && shiftr_good rshiftr2 ry offset2)
+ (pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ') *' pcst (#pattern.ident.Z_shiftr @ ??' @ #?ℤ'))
+ (fun r rshiftr1 rx x roffset1 offset1 rshiftr2 ry y roffset2 offset2 => let s := (2*offset1)%Z in cst r (#(ident.fancy_mulhh s) @ (cst rx x, cst ry y)) when (offset1 =? offset2) && shiftr_good rshiftr1 rx offset1 && shiftr_good rshiftr2 ry offset2 && lit_good offset1 roffset1 && lit_good offset2 roffset2)
@@ -2595,12 +2458,29 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y)
Redirect "arith_with_casts_rewrite_head" Print arith_with_casts_rewrite_head.
End red_arith_with_casts.
+ Section red_strip_literal_casts.
+ Context {var : type.type base.type -> Type}
+ (do_again : forall t : base.type, @expr base.type ident (@Compile.value base.type ident var) (type.base t)
+ -> @UnderLets.UnderLets base.type ident var (@expr base.type ident var (type.base t)))
+ {t} (idc : ident t).
+
+ Time Definition strip_literal_casts_rewrite_head
+ := make_rewrite_head (@strip_literal_casts_rewrite_head0 var do_again t idc) (@strip_literal_casts_pr2_rewrite_rules).
+
+ Local Set Printing Depth 1000000.
+ Local Set Printing Width 200.
+ Import RewriterPrintingNotations.
+ Redirect "strip_literal_casts_rewrite_head" Print strip_literal_casts_rewrite_head.
+ End red_strip_literal_casts.
+
Definition RewriteNBE {t} (e : expr.Expr (ident:=ident) t) : expr.Expr (ident:=ident) t
:= @Compile.Rewrite (@nbe_rewrite_head) nbe_default_fuel t e.
Definition RewriteArith (max_const_val : Z) {t} (e : expr.Expr (ident:=ident) t) : expr.Expr (ident:=ident) t
:= @Compile.Rewrite (@arith_rewrite_head max_const_val) arith_default_fuel t e.
Definition RewriteArithWithCasts {t} (e : expr.Expr (ident:=ident) t) : expr.Expr (ident:=ident) t
:= @Compile.Rewrite (@arith_with_casts_rewrite_head) arith_with_casts_default_fuel t e.
+ Definition RewriteStripLiteralCasts {t} (e : expr.Expr (ident:=ident) t) : expr.Expr (ident:=ident) t
+ := @Compile.Rewrite (fun var do_again => @strip_literal_casts_rewrite_head var) strip_literal_casts_default_fuel t e.
Definition RewriteToFancy
(invert_low invert_high : Z (*log2wordmax*) -> Z -> @option Z)
{t} (e : expr.Expr (ident:=ident) t) : expr.Expr (ident:=ident) t
diff --git a/src/RewriterProofs.v b/src/RewriterProofs.v
index c18527305..e4ac4afda 100644
--- a/src/RewriterProofs.v
+++ b/src/RewriterProofs.v
@@ -94,6 +94,11 @@ Module Compilers.
start_Wf_proof arith_with_casts_rewrite_head_eq arith_with_casts_all_rewrite_rules_eq (@arith_with_casts_rewrite_head0).
apply arith_with_casts_rewrite_rules_good.
Qed.
+ Lemma Wf_RewriteStripLiteralCasts {t} e (Hwf : Wf e) : Wf (@RewriteStripLiteralCasts t e).
+ Proof.
+ start_Wf_proof strip_literal_casts_rewrite_head_eq strip_literal_casts_all_rewrite_rules_eq (@strip_literal_casts_rewrite_head0).
+ apply strip_literal_casts_rewrite_rules_good.
+ Qed.
Lemma Wf_RewriteToFancy (invert_low invert_high : Z -> Z -> option Z)
(Hlow : forall s v v', invert_low s v = Some v' -> v = Z.land v' (2^(s/2)-1))
(Hhigh : forall s v v', invert_high s v = Some v' -> v = Z.shiftr v' (s/2))
@@ -136,6 +141,14 @@ Module Compilers.
apply arith_with_casts_rewrite_rules_interp_good.
Qed.
+ Lemma Interp_gen_RewriteStripLiteralCasts {cast_outside_of_range} {t} e (Hwf : Wf e)
+ : expr.Interp (@ident.gen_interp cast_outside_of_range) (@RewriteStripLiteralCasts t e)
+ == expr.Interp (@ident.gen_interp cast_outside_of_range) e.
+ Proof.
+ start_Interp_proof strip_literal_casts_rewrite_head_eq strip_literal_casts_all_rewrite_rules_eq (@strip_literal_casts_rewrite_head0).
+ apply strip_literal_casts_rewrite_rules_interp_good.
+ Qed.
+
Lemma Interp_gen_RewriteToFancy {cast_outside_of_range} (invert_low invert_high : Z -> Z -> option Z)
(Hlow : forall s v v', invert_low s v = Some v' -> v = Z.land v' (2^(s/2)-1))
(Hhigh : forall s v v', invert_high s v = Some v' -> v = Z.shiftr v' (s/2))
@@ -167,6 +180,9 @@ Module Compilers.
Lemma Interp_RewriteArithWithCasts {t} e (Hwf : Wf e)
: Interp (@RewriteArithWithCasts t e) == Interp e.
Proof. apply Interp_gen_RewriteArithWithCasts; assumption. Qed.
+ Lemma Interp_RewriteStripLiteralCasts {t} e (Hwf : Wf e)
+ : Interp (@RewriteStripLiteralCasts t e) == Interp e.
+ Proof. apply Interp_gen_RewriteStripLiteralCasts; assumption. Qed.
Lemma Interp_RewriteToFancy (invert_low invert_high : Z -> Z -> option Z)
(Hlow : forall s v v', invert_low s v = Some v' -> v = Z.land v' (2^(s/2)-1))
(Hhigh : forall s v v', invert_high s v = Some v' -> v = Z.shiftr v' (s/2))
@@ -193,6 +209,6 @@ Module Compilers.
: Interp (@PartialEvaluate t e) == Interp e.
Proof. apply Interp_gen_PartialEvaluate; assumption. Qed.
- Hint Resolve Wf_PartialEvaluate Wf_RewriteArith Wf_RewriteNBE Wf_RewriteToFancy Wf_RewriteArithWithCasts Wf_RewriteToFancyWithCasts : wf.
- Hint Rewrite @Interp_gen_PartialEvaluate @Interp_gen_RewriteArith @Interp_gen_RewriteNBE @Interp_gen_RewriteToFancy @Interp_gen_RewriteArithWithCasts @Interp_gen_RewriteToFancyWithCasts @Interp_PartialEvaluate @Interp_RewriteArith @Interp_RewriteNBE @Interp_RewriteToFancy @Interp_RewriteArithWithCasts @Interp_RewriteToFancyWithCasts : interp.
+ Hint Resolve Wf_PartialEvaluate Wf_RewriteArith Wf_RewriteNBE Wf_RewriteToFancy Wf_RewriteArithWithCasts Wf_RewriteStripLiteralCasts Wf_RewriteToFancyWithCasts : wf.
+ Hint Rewrite @Interp_gen_PartialEvaluate @Interp_gen_RewriteArith @Interp_gen_RewriteNBE @Interp_gen_RewriteToFancy @Interp_gen_RewriteArithWithCasts @Interp_gen_RewriteStripLiteralCasts @Interp_gen_RewriteToFancyWithCasts @Interp_PartialEvaluate @Interp_RewriteArith @Interp_RewriteNBE @Interp_RewriteToFancy @Interp_RewriteArithWithCasts @Interp_RewriteStripLiteralCasts @Interp_RewriteToFancyWithCasts : interp.
End Compilers.
diff --git a/src/RewriterRulesGood.v b/src/RewriterRulesGood.v
index 3718f04e5..9deb2ec6d 100644
--- a/src/RewriterRulesGood.v
+++ b/src/RewriterRulesGood.v
@@ -62,6 +62,9 @@ Module Compilers.
Lemma arith_with_casts_rewrite_head_eq : @arith_with_casts_rewrite_head = @arith_with_casts_rewrite_head0.
Proof. reflexivity. Qed.
+ Lemma strip_literal_casts_rewrite_head_eq : (fun var do_again => @strip_literal_casts_rewrite_head var) = @strip_literal_casts_rewrite_head0.
+ Proof. reflexivity. Qed.
+
Lemma nbe_all_rewrite_rules_eq : @nbe_all_rewrite_rules = @nbe_rewrite_rules.
Proof. reflexivity. Qed.
@@ -77,6 +80,9 @@ Module Compilers.
Lemma arith_with_casts_all_rewrite_rules_eq : @arith_with_casts_all_rewrite_rules = @arith_with_casts_rewrite_rules.
Proof. reflexivity. Qed.
+ Lemma strip_literal_casts_all_rewrite_rules_eq : @strip_literal_casts_all_rewrite_rules = @strip_literal_casts_rewrite_rules.
+ Proof. reflexivity. Qed.
+
Section good.
Context {var1 var2 : type -> Type}.
@@ -358,6 +364,13 @@ Module Compilers.
Time all: repeat good_t_step.
Qed.
+ Lemma strip_literal_casts_rewrite_rules_good
+ : rewrite_rules_goodT strip_literal_casts_rewrite_rules strip_literal_casts_rewrite_rules.
+ Proof using Type.
+ Time start_good.
+ Time all: repeat good_t_step.
+ Qed.
+
Lemma fancy_rewrite_rules_good
(invert_low invert_high : Z -> Z -> option Z)
(Hlow : forall s v v', invert_low s v = Some v' -> v = Z.land v' (2^(s/2)-1))
diff --git a/src/RewriterRulesInterpGood.v b/src/RewriterRulesInterpGood.v
index 469f0cb31..612f72afd 100644
--- a/src/RewriterRulesInterpGood.v
+++ b/src/RewriterRulesInterpGood.v
@@ -743,6 +743,13 @@ Module Compilers.
Time all: try solve [ repeat interp_good_t_step_related; repeat interp_good_t_step_arith; fin_with_nia ].
Qed.
+ Lemma strip_literal_casts_rewrite_rules_interp_good
+ : rewrite_rules_interp_goodT strip_literal_casts_rewrite_rules.
+ Proof using Type.
+ Time start_interp_good.
+ Time all: try solve [ repeat interp_good_t_step_related; repeat interp_good_t_step_arith ].
+ Qed.
+
Local Ltac fancy_local_t :=
repeat match goal with
| [ H : forall s v v', ?invert_low s v = Some v' -> v = _,
diff --git a/src/arith_with_casts_rewrite_head.out b/src/arith_with_casts_rewrite_head.out
index 365a351b0..b7bec9985 100644
--- a/src/arith_with_casts_rewrite_head.out
+++ b/src/arith_with_casts_rewrite_head.out
@@ -216,77 +216,150 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
| @List_nth_default T =>
fun (x : expr T) (x0 : expr (list T)) (x1 : expr ℕ) =>
Base (#(List_nth_default)%expr @ x @ x0 @ x1)%expr_pat
-| Z_add => fun x x0 : expr ℤ => Base (x + x0)%expr
+| Z_add =>
+ fun x x0 : expr ℤ =>
+ (((match x with
+ | (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat =>
+ args <- invert_bind_args idc0 Raw.ident.Literal;
+ args0 <- invert_bind_args idc Raw.ident.Z_cast;
+ match
+ pattern.type.unify_extracted (ℤ -> ℤ)%ptype
+ ((projT1 args) -> ℤ)%ptype
+ with
+ | Some (_, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ ((projT1 args) -> ℤ)%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal ##(projT2 args);
+ fv <- (x2 <- (if
+ ((let (x2, _) := xv in x2) =? 0) &&
+ is_bounded_by_bool (let (x2, _) := xv in x2)
+ (ZRange.normalize args0)
+ then Some x0
+ else None);
+ Some (Base x2));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | (@expr.Ident _ _ _ t idc @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
+ | _ => None
+ end;;
+ match x0 with
+ | (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat =>
+ args <- invert_bind_args idc0 Raw.ident.Literal;
+ args0 <- invert_bind_args idc Raw.ident.Z_cast;
+ match
+ pattern.type.unify_extracted (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ with
+ | Some (_, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
+ (ℤ -> (projT1 args))%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal ##(projT2 args);
+ fv <- (x2 <- (if
+ ((let (x2, _) := xv in x2) =? 0) &&
+ is_bounded_by_bool (let (x2, _) := xv in x2)
+ (ZRange.normalize args0)
+ then Some x
+ else None);
+ Some (Base x2));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | (@expr.Ident _ _ _ t idc @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
+ | _ => None
+ end);;
+ None);;;
+ Base (x + x0)%expr)%option
| Z_mul => fun x x0 : expr ℤ => Base (x * x0)%expr
| Z_pow => fun x x0 : expr ℤ => Base (#(Z_pow)%expr @ x @ x0)%expr_pat
| Z_sub =>
fun x x0 : expr ℤ =>
((match x with
- | @expr.Ident _ _ _ t idc =>
+ | (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat =>
match x0 with
- | (@expr.Ident _ _ _ t0 idc0 @
- (@expr.Ident _ _ _ t1 idc1 @ @expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t2 idc2) x3))%expr_pat =>
- args <- invert_bind_args idc2 Raw.ident.Z_cast;
- _ <- invert_bind_args idc1 Raw.ident.Z_opp;
- args1 <- invert_bind_args idc0 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc Raw.ident.Literal;
+ | (@expr.Ident _ _ _ t1 idc1 @
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s2 _
+ (@expr.Ident _ _ _ t3 idc3) x4))%expr_pat =>
+ args <- invert_bind_args idc3 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc2 Raw.ident.Z_opp;
+ args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc0 Raw.ident.Literal;
+ args3 <- invert_bind_args idc Raw.ident.Z_cast;
match
pattern.type.unify_extracted (ℤ -> ℤ)%ptype
- ((projT1 args2) -> s1)%ptype
+ ((projT1 args2) -> s2)%ptype
with
| Some (_, _)%zrange =>
if
type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
- ((projT1 args2) -> s1)%ptype
+ ((projT1 args2) -> s2)%ptype
then
xv <- ident.unify pattern.ident.Literal ##(projT2 args2);
- v <- type.try_make_transport_cps s1 ℤ;
- fv <- (x4 <- (if
- ((let (x4, _) := xv in x4) =? 0) &&
+ v <- type.try_make_transport_cps s2 ℤ;
+ fv <- (x5 <- (if
+ ((let (x5, _) := xv in x5) =? 0) &&
(ZRange.normalize args <=?
- - ZRange.normalize args1)%zrange
+ - ZRange.normalize args1)%zrange &&
+ is_bounded_by_bool
+ (let (x5, _) := xv in x5) args3
then
Some
(#(Z_cast args)%expr @
- v (Compile.reflect x3))%expr_pat
+ v (Compile.reflect x4))%expr_pat
else None);
- Some (Base x4));
+ Some (Base x5));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | (@expr.Ident _ _ _ t0 idc0 @
- (@expr.Ident _ _ _ t1 idc1 @ @expr.App _ _ _ s1 _ ($_)%expr _))%expr_pat |
- (@expr.Ident _ _ _ t0 idc0 @
- (@expr.Ident _ _ _ t1 idc1 @ @expr.App _ _ _ s1 _
+ | (@expr.Ident _ _ _ t1 idc1 @
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s2 _ ($_)%expr _))%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s2 _
(@expr.Abs _ _ _ _ _ _) _))%expr_pat |
- (@expr.Ident _ _ _ t0 idc0 @
- (@expr.Ident _ _ _ t1 idc1 @ @expr.App _ _ _ s1 _ (_ @ _) _))%expr_pat |
- (@expr.Ident _ _ _ t0 idc0 @
- (@expr.Ident _ _ _ t1 idc1 @ @expr.App _ _ _ s1 _
+ (@expr.Ident _ _ _ t1 idc1 @
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s2 _ (_ @ _) _))%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s2 _
(@expr.LetIn _ _ _ _ _ _ _) _))%expr_pat => None
- | (@expr.Ident _ _ _ t0 idc0 @ (@expr.Ident _ _ _ t1 idc1 @ #(_)))%expr_pat |
- (@expr.Ident _ _ _ t0 idc0 @
- (@expr.Ident _ _ _ t1 idc1 @ ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t0 idc0 @
- (@expr.Ident _ _ _ t1 idc1 @ @expr.Abs _ _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t0 idc0 @
- (@expr.Ident _ _ _ t1 idc1 @ @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
+ | (@expr.Ident _ _ _ t1 idc1 @ (@expr.Ident _ _ _ t2 idc2 @ #(_)))%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @
+ (@expr.Ident _ _ _ t2 idc2 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.Abs _ _ _ _ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
None
- | (@expr.Ident _ _ _ t0 idc0 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t0 idc0 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t0 idc0 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t0 idc0 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t0 idc0 @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t0 idc0 @ (_ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t0 idc0 @ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t0 idc0 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ | (@expr.Ident _ _ _ t1 idc1 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @ (_ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
None
| _ => None
end;;
- args <- invert_bind_args idc Raw.ident.Literal;
+ args <- invert_bind_args idc0 Raw.ident.Literal;
+ args0 <- invert_bind_args idc Raw.ident.Z_cast;
match
pattern.type.unify_extracted (ℤ -> ℤ)%ptype
((projT1 args) -> ℤ)%ptype
@@ -297,15 +370,23 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> ℤ)%ptype
then
xv <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (x1 <- (if (let (x1, _) := xv in x1) =? 0
+ fv <- (x2 <- (if
+ ((let (x2, _) := xv in x2) =? 0) &&
+ is_bounded_by_bool (let (x2, _) := xv in x2)
+ (ZRange.normalize args0)
then Some (- x0)%expr
else None);
- Some (Base x1));
+ Some (Base x2));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
+ | (@expr.Ident _ _ _ t idc @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end;;
None);;;
@@ -380,8 +461,9 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
| Z_shiftl =>
fun x x0 : expr ℤ =>
(match x with
- | @expr.Ident _ _ _ t idc =>
- args <- invert_bind_args idc Raw.ident.Literal;
+ | (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat =>
+ args <- invert_bind_args idc0 Raw.ident.Literal;
+ args0 <- invert_bind_args idc Raw.ident.Z_cast;
match
pattern.type.unify_extracted (ℤ -> ℤ)%ptype
((projT1 args) -> ℤ)%ptype
@@ -392,15 +474,22 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((projT1 args) -> ℤ)%ptype
then
xv <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (x1 <- (if (let (x1, _) := xv in x1) =? 0
+ fv <- (x2 <- (if
+ ((let (x2, _) := xv in x2) =? 0) &&
+ is_bounded_by_bool (let (x2, _) := xv in x2)
+ (ZRange.normalize args0)
then Some (##0)%expr
else None);
- Some (Base x1));
+ Some (Base x2));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
+ | (@expr.Ident _ _ _ t idc @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat => None
| _ => None
end;;;
Base (x << x0)%expr)%option
@@ -411,269 +500,375 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
fun x x0 : expr ℤ => Base (#(Z_lnot_modulo)%expr @ x @ x0)%expr_pat
| Z_mul_split =>
fun x x0 x1 : expr ℤ =>
- ((match x with
- | @expr.Ident _ _ _ t idc =>
- match x0 with
- | @expr.Ident _ _ _ t0 idc0 =>
- args <- invert_bind_args idc0 Raw.ident.Literal;
- args0 <- invert_bind_args idc Raw.ident.Literal;
- match
- pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args0) -> (projT1 args)) -> ℤ)%ptype
- with
- | Some (_, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args0) -> (projT1 args)) -> ℤ)%ptype
- then
- _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
- xv0 <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (x2 <- (if (let (x2, _) := xv0 in x2) =? 0
- then Some ((##0)%expr, (##0)%expr)%expr_pat
- else None);
- Some (Base x2));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end;;
- match x1 with
- | @expr.Ident _ _ _ t0 idc0 =>
- (args <- invert_bind_args idc0 Raw.ident.Literal;
- args0 <- invert_bind_args idc Raw.ident.Literal;
+ (((match x0 with
+ | (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat =>
+ args <- invert_bind_args idc0 Raw.ident.Literal;
+ args0 <- invert_bind_args idc Raw.ident.Z_cast;
+ match
+ pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
+ ((ℤ -> (projT1 args)) -> ℤ)%ptype
+ with
+ | Some (_, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype ((ℤ -> (projT1 args)) -> ℤ)%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal ##(projT2 args);
+ fv <- (x3 <- (if
+ ((let (x3, _) := xv in x3) =? 0) &&
+ is_bounded_by_bool (let (x3, _) := xv in x3)
+ (ZRange.normalize args0)
+ then
+ Some
+ (#(Z_cast r[0 ~> 0])%expr @ (##0)%expr,
+ #(Z_cast r[0 ~> 0])%expr @ (##0)%expr)%expr_pat
+ else None);
+ Some (Base x3));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | (@expr.Ident _ _ _ t idc @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
+ | _ => None
+ end;;
+ match x1 with
+ | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x2 =>
+ match x2 with
+ | @expr.Ident _ _ _ t0 idc0 =>
+ args <- invert_bind_args idc0 Raw.ident.Literal;
+ args0 <- invert_bind_args idc Raw.ident.Z_cast;
match
pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args0) -> ℤ) -> (projT1 args))%ptype
+ ((ℤ -> ℤ) -> (projT1 args))%ptype
with
| Some (_, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args0) -> ℤ) -> (projT1 args))%ptype
+ ((ℤ -> ℤ) -> ℤ)%ptype ((ℤ -> ℤ) -> (projT1 args))%ptype
then
- _ <- ident.unify pattern.ident.Literal ##(projT2 args0);
- xv0 <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (x2 <- (if (let (x2, _) := xv0 in x2) =? 0
- then Some ((##0)%expr, (##0)%expr)%expr_pat
+ xv <- ident.unify pattern.ident.Literal ##(projT2 args);
+ fv <- (x3 <- (if
+ ((let (x3, _) := xv in x3) =? 0) &&
+ is_bounded_by_bool
+ (let (x3, _) := xv in x3)
+ (ZRange.normalize args0)
+ then
+ Some
+ (#(Z_cast r[0 ~> 0])%expr @ (##0)%expr,
+ #(Z_cast r[0 ~> 0])%expr @ (##0)%expr)%expr_pat
else None);
- Some (Base x2));
+ Some (Base x3));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
- end);;
- match x0 with
- | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t1 idc1) x2 =>
- args <- invert_bind_args idc1 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc0 Raw.ident.Literal;
- args1 <- invert_bind_args idc Raw.ident.Literal;
- match
- pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s) -> (projT1 args0))%ptype
- with
- | Some (_, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s) -> (projT1 args0))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- fv <- (x3 <- (if
- ((let (x3, _) := xv0 in x3) =? 1) &&
- (ZRange.normalize args <=?
- r[0 ~> (let (x3, _) := xv in x3) - 1])%zrange
- then
- Some
- (#(Z_cast args)%expr @
- v (Compile.reflect x2), (##0)%expr)%expr_pat
- else None);
- Some (Base x3));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t0 idc0) x2 =>
- match x0 with
- | @expr.Ident _ _ _ t1 idc1 =>
- args <- invert_bind_args idc1 Raw.ident.Literal;
- args0 <- invert_bind_args idc0 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc Raw.ident.Literal;
- match
- pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> (projT1 args)) -> s)%ptype
- with
- | Some (_, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> (projT1 args)) -> s)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v <- type.try_make_transport_cps s ℤ;
- fv <- (x3 <- (if
- ((let (x3, _) := xv0 in x3) =? 1) &&
- (ZRange.normalize args0 <=?
- r[0 ~> (let (x3, _) := xv in x3) - 1])%zrange
- then
- Some
- (#(Z_cast args0)%expr @
- v (Compile.reflect x2), (##0)%expr)%expr_pat
- else None);
- Some (Base x3));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat
- _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | _ => None
- end;;
+ end
+ | _ => None
+ end;;
+ match x with
+ | (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat =>
+ match x0 with
+ | @expr.App _ _ _ s1 _ (@expr.Ident _ _ _ t2 idc2) x4 =>
+ match x4 with
+ | @expr.Ident _ _ _ t3 idc3 =>
+ args <- invert_bind_args idc3 Raw.ident.Literal;
+ args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc1 Raw.ident.Literal;
+ args2 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc Raw.ident.Z_cast;
+ match
+ pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> (projT1 args)) -> s)%ptype
+ with
+ | Some (_, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> (projT1 args)) -> s)%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v <- type.try_make_transport_cps s ℤ;
+ fv <- (x5 <- (if
+ ((let (x5, _) := xv0 in x5) =? 1) &&
+ (ZRange.normalize args3 <=?
+ r[0 ~> (let (x5, _) := xv in x5) -
+ 1])%zrange &&
+ is_bounded_by_bool
+ (let (x5, _) := xv in x5)
+ (ZRange.normalize args2) &&
+ is_bounded_by_bool
+ (let (x5, _) := xv0 in x5)
+ (ZRange.normalize args0)
+ then
+ Some
+ (#(Z_cast args3)%expr @
+ v (Compile.reflect x2),
+ #(Z_cast r[0 ~> 0])%expr @
+ (##0)%expr)%expr_pat
+ else None);
+ Some (Base x5));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | _ => None
+ end;;
+ match x2 with
+ | @expr.Ident _ _ _ t3 idc3 =>
+ args <- invert_bind_args idc3 Raw.ident.Literal;
+ args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc1 Raw.ident.Literal;
+ args2 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc Raw.ident.Z_cast;
+ match
+ pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> s1) -> (projT1 args))%ptype
+ with
+ | Some (_, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> s1) -> (projT1 args))%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s1 ℤ;
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (x5 <- (if
+ ((let (x5, _) := xv0 in x5) =? 1) &&
+ (ZRange.normalize args0 <=?
+ r[0 ~> (let (x5, _) := xv in x5) -
+ 1])%zrange &&
+ is_bounded_by_bool
+ (let (x5, _) := xv in x5)
+ (ZRange.normalize args2) &&
+ is_bounded_by_bool
+ (let (x5, _) := xv0 in x5)
+ (ZRange.normalize args3)
+ then
+ Some
+ (#(Z_cast args0)%expr @
+ v (Compile.reflect x4),
+ #(Z_cast r[0 ~> 0])%expr @
+ (##0)%expr)%expr_pat
+ else None);
+ Some (Base x5));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | _ => None
+ end
+ | @expr.App _ _ _ s1 _ ($_)%expr _ | @expr.App _ _ _ s1 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s1 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s1 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end
+ | (@expr.Ident _ _ _ t0 idc0 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
+ | _ => None
+ end
+ | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end);;
None);;;
Base (#(Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat)%option
| Z_add_get_carry =>
fun x x0 x1 : expr ℤ =>
((match x with
- | @expr.Ident _ _ _ t idc =>
+ | (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat =>
match x0 with
- | @expr.Ident _ _ _ t0 idc0 =>
- match x1 with
- | @expr.Ident _ _ _ t1 idc1 =>
- args <- invert_bind_args idc1 Raw.ident.Literal;
- args0 <- invert_bind_args idc0 Raw.ident.Literal;
- args1 <- invert_bind_args idc Raw.ident.Literal;
- match
- pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> (projT1 args0)) -> (projT1 args))%ptype
- with
- | Some (_, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> (projT1 args0)) -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (let
- '(a1, b1)%zrange :=
- Z.add_get_carry_full (let (x2, _) := xv in x2)
- (let (x2, _) := xv0 in x2)
- (let (x2, _) := xv1 in x2) in
- ((##a1)%expr, (##b1)%expr)%expr_pat))
- else None
- | None => None
- end
- | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t1 idc1) x2 =>
- args <- invert_bind_args idc1 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc0 Raw.ident.Literal;
- args1 <- invert_bind_args idc Raw.ident.Literal;
- match
- pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> (projT1 args0)) -> s)%ptype
- with
- | Some (_, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> (projT1 args0)) -> s)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s ℤ;
- fv <- (x3 <- (if
- ((let (x3, _) := xv0 in x3) =? 0) &&
- (ZRange.normalize args <=?
- r[0 ~> (let (x3, _) := xv in x3) - 1])%zrange
- then
- Some
- (#(Z_cast args)%expr @
- v (Compile.reflect x2), (##0)%expr)%expr_pat
- else None);
- Some (Base x3));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
+ | @expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t1 idc1) x3 =>
+ match x3 with
+ | @expr.Ident _ _ _ t2 idc2 =>
+ match x1 with
+ | @expr.App _ _ _ s1 _ (@expr.Ident _ _ _ t3 idc3) x4 =>
+ match x4 with
+ | @expr.Ident _ _ _ t4 idc4 =>
+ args <- invert_bind_args idc4 Raw.ident.Literal;
+ args0 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc2 Raw.ident.Literal;
+ args2 <- invert_bind_args idc1 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc0 Raw.ident.Literal;
+ args4 <- invert_bind_args idc Raw.ident.Z_cast;
+ match
+ pattern.type.unify_extracted
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args3) -> (projT1 args1)) ->
+ (projT1 args))%ptype
+ with
+ | Some (_, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args3) -> (projT1 args1)) ->
+ (projT1 args))%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args3);
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ xv1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (x5 <- (if
+ is_bounded_by_bool
+ (let (x5, _) := xv in x5)
+ (ZRange.normalize args4) &&
+ is_bounded_by_bool
+ (let (x5, _) := xv0 in x5)
+ (ZRange.normalize args2) &&
+ is_bounded_by_bool
+ (let (x5, _) := xv1 in x5)
+ (ZRange.normalize args0)
+ then
+ Some
+ (let
+ '(a1, b1)%zrange :=
+ Z.add_get_carry_full
+ (let (x5, _) := xv in x5)
+ (let (x5, _) := xv0 in x5)
+ (let (x5, _) := xv1 in x5)
+ in
+ ((##a1)%expr, (##b1)%expr)%expr_pat)
+ else None);
+ Some (Base x5));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | _ => None
+ end;;
+ args <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args0 <- invert_bind_args idc2 Raw.ident.Literal;
+ args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc0 Raw.ident.Literal;
+ args3 <- invert_bind_args idc Raw.ident.Z_cast;
+ match
+ pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args2) -> (projT1 args0)) -> s1)%ptype
+ with
+ | Some (_, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args2) -> (projT1 args0)) -> s1)%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args2);
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ v <- type.try_make_transport_cps s1 ℤ;
+ fv <- (x5 <- (if
+ ((let (x5, _) := xv0 in x5) =? 0) &&
+ (ZRange.normalize args <=?
+ r[0 ~> (let (x5, _) := xv in x5) -
+ 1])%zrange &&
+ is_bounded_by_bool
+ (let (x5, _) := xv0 in x5)
+ (ZRange.normalize args1) &&
+ is_bounded_by_bool
+ (let (x5, _) := xv in x5)
+ (ZRange.normalize args3)
+ then
+ Some
+ (#(Z_cast args)%expr @
+ v (Compile.reflect x4),
+ #(Z_cast r[0 ~> 0])%expr @
+ (##0)%expr)%expr_pat
+ else None);
+ Some (Base x5));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | @expr.App _ _ _ s1 _ ($_)%expr _ | @expr.App _ _ _ s1 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s1 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s1 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
- end
- | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t0 idc0) x2 =>
+ end;;
match x1 with
- | @expr.Ident _ _ _ t1 idc1 =>
- args <- invert_bind_args idc1 Raw.ident.Literal;
- args0 <- invert_bind_args idc0 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc Raw.ident.Literal;
+ | (@expr.Ident _ _ _ t2 idc2 @ @expr.Ident _ _ _ t3 idc3)%expr_pat =>
+ args <- invert_bind_args idc3 Raw.ident.Literal;
+ args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc0 Raw.ident.Literal;
+ args3 <- invert_bind_args idc Raw.ident.Z_cast;
match
pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s) -> (projT1 args))%ptype
+ (((projT1 args2) -> s0) -> (projT1 args))%ptype
with
| Some (_, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s) -> (projT1 args))%ptype
+ (((projT1 args2) -> s0) -> (projT1 args))%ptype
then
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s ℤ;
+ ##(projT2 args2);
+ v <- type.try_make_transport_cps s0 ℤ;
xv0 <- ident.unify pattern.ident.Literal
##(projT2 args);
- fv <- (x3 <- (if
- ((let (x3, _) := xv0 in x3) =? 0) &&
- (ZRange.normalize args0 <=?
- r[0 ~> (let (x3, _) := xv in x3) - 1])%zrange
+ fv <- (x5 <- (if
+ ((let (x5, _) := xv0 in x5) =? 0) &&
+ (ZRange.normalize args1 <=?
+ r[0 ~> (let (x5, _) := xv in x5) - 1])%zrange &&
+ is_bounded_by_bool
+ (let (x5, _) := xv0 in x5)
+ (ZRange.normalize args0) &&
+ is_bounded_by_bool
+ (let (x5, _) := xv in x5)
+ (ZRange.normalize args3)
then
Some
- (#(Z_cast args0)%expr @
- v (Compile.reflect x2), (##0)%expr)%expr_pat
+ (#(Z_cast args1)%expr @
+ v (Compile.reflect x3),
+ #(Z_cast r[0 ~> 0])%expr @ (##0)%expr)%expr_pat
else None);
- Some (Base x3));
+ Some (Base x5));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
+ | (@expr.Ident _ _ _ t2 idc2 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _ (_ @ _)%expr_pat
- _ | @expr.App _ _ _ s _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | @expr.App _ _ _ s0 _ ($_)%expr _ | @expr.App _ _ _ s0 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s0 _ (_ @ _)%expr_pat
+ _ | @expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
+ | (@expr.Ident _ _ _ t idc @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end;;
None);;;
@@ -681,8 +876,9 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
| Z_add_with_carry =>
fun x x0 x1 : expr ℤ =>
(match x with
- | @expr.Ident _ _ _ t idc =>
- args <- invert_bind_args idc Raw.ident.Literal;
+ | (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat =>
+ args <- invert_bind_args idc0 Raw.ident.Literal;
+ args0 <- invert_bind_args idc Raw.ident.Z_cast;
match
pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
(((projT1 args) -> ℤ) -> ℤ)%ptype
@@ -693,170 +889,249 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((ℤ -> ℤ) -> ℤ)%ptype (((projT1 args) -> ℤ) -> ℤ)%ptype
then
xv <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (x2 <- (if (let (x2, _) := xv in x2) =? 0
+ fv <- (x3 <- (if
+ ((let (x3, _) := xv in x3) =? 0) &&
+ is_bounded_by_bool (let (x3, _) := xv in x3)
+ (ZRange.normalize args0)
then Some (x0 + x1)%expr
else None);
- Some (Base x2));
+ Some (Base x3));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
+ | (@expr.Ident _ _ _ t idc @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat => None
| _ => None
end;;;
Base (#(Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat)%option
| Z_add_with_get_carry =>
fun x x0 x1 x2 : expr ℤ =>
(match x with
- | @expr.Ident _ _ _ t idc =>
+ | (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat =>
match x0 with
- | @expr.Ident _ _ _ t0 idc0 =>
+ | (@expr.Ident _ _ _ t1 idc1 @ @expr.Ident _ _ _ t2 idc2)%expr_pat =>
match x1 with
- | @expr.Ident _ _ _ t1 idc1 =>
- match x2 with
- | @expr.Ident _ _ _ t2 idc2 =>
- args <- invert_bind_args idc2 Raw.ident.Literal;
- args0 <- invert_bind_args idc1 Raw.ident.Literal;
- args1 <- invert_bind_args idc0 Raw.ident.Literal;
- args2 <- invert_bind_args idc Raw.ident.Literal;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) ->
- (projT1 args0)) -> (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) ->
- (projT1 args0)) -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- xv2 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- Some
- (Base
- (let
- '(a2, b2)%zrange :=
- Z.add_with_get_carry_full
- (let (x3, _) := xv in x3)
- (let (x3, _) := xv0 in x3)
- (let (x3, _) := xv1 in x3)
- (let (x3, _) := xv2 in x3) in
- ((##a2)%expr, (##b2)%expr)%expr_pat))
- else None
- | None => None
- end
- | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t2 idc2) x3 =>
- args <- invert_bind_args idc2 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc1 Raw.ident.Literal;
- args1 <- invert_bind_args idc0 Raw.ident.Literal;
- args2 <- invert_bind_args idc Raw.ident.Literal;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) ->
- (projT1 args0)) -> s)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) ->
- (projT1 args0)) -> s)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s ℤ;
- fv <- (x4 <- (if
- ((let (x4, _) := xv0 in x4) =? 0) &&
- ((let (x4, _) := xv1 in x4) =? 0) &&
- (ZRange.normalize args <=?
- r[0 ~> (let (x4, _) := xv in x4) -
- 1])%zrange
- then
- Some
- (#(Z_cast args)%expr @
- v (Compile.reflect x3),
- (##0)%expr)%expr_pat
- else None);
- Some (Base x4));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
+ | @expr.App _ _ _ s1 _ (@expr.Ident _ _ _ t3 idc3) x5 =>
+ match x5 with
+ | @expr.Ident _ _ _ t4 idc4 =>
+ match x2 with
+ | @expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t5 idc5) x6 =>
+ match x6 with
+ | @expr.Ident _ _ _ t6 idc6 =>
+ args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Literal;
+ args2 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc2 Raw.ident.Literal;
+ args4 <- invert_bind_args idc1 Raw.ident.Z_cast;
+ args5 <- invert_bind_args idc0 Raw.ident.Literal;
+ args6 <- invert_bind_args idc Raw.ident.Z_cast;
+ match
+ pattern.type.unify_extracted
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args5) -> (projT1 args3)) ->
+ (projT1 args1)) -> (projT1 args))%ptype
+ with
+ | Some (_, _, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args5) -> (projT1 args3)) ->
+ (projT1 args1)) -> (projT1 args))%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args5);
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args3);
+ xv1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ xv2 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (x7 <- (if
+ is_bounded_by_bool
+ (let (x7, _) := xv in x7)
+ (ZRange.normalize args6) &&
+ is_bounded_by_bool
+ (let (x7, _) := xv0 in x7)
+ (ZRange.normalize args4) &&
+ is_bounded_by_bool
+ (let (x7, _) := xv1 in x7)
+ (ZRange.normalize args2) &&
+ is_bounded_by_bool
+ (let (x7, _) := xv2 in x7)
+ (ZRange.normalize args0)
+ then
+ Some
+ (let
+ '(a2, b2)%zrange :=
+ Z.add_with_get_carry_full
+ (let (x7, _) := xv in
+ x7)
+ (let (x7, _) := xv0 in
+ x7)
+ (let (x7, _) := xv1 in
+ x7)
+ (let (x7, _) := xv2 in
+ x7) in
+ ((##a2)%expr,
+ (##b2)%expr)%expr_pat)
+ else None);
+ Some (Base x7));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | _ => None
+ end;;
+ args <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args0 <- invert_bind_args idc4 Raw.ident.Literal;
+ args1 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc2 Raw.ident.Literal;
+ args3 <- invert_bind_args idc1 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc0 Raw.ident.Literal;
+ args5 <- invert_bind_args idc Raw.ident.Z_cast;
+ match
+ pattern.type.unify_extracted
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args4) -> (projT1 args2)) ->
+ (projT1 args0)) -> s2)%ptype
+ with
+ | Some (_, _, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args4) -> (projT1 args2)) ->
+ (projT1 args0)) -> s2)%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args4);
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args2);
+ xv1 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ v <- type.try_make_transport_cps s2 ℤ;
+ fv <- (x7 <- (if
+ ((let (x7, _) := xv0 in x7) =? 0) &&
+ ((let (x7, _) := xv1 in x7) =? 0) &&
+ (ZRange.normalize args <=?
+ r[0 ~> (let (x7, _) := xv in x7) -
+ 1])%zrange &&
+ is_bounded_by_bool
+ (let (x7, _) := xv in x7)
+ (ZRange.normalize args5) &&
+ is_bounded_by_bool
+ (let (x7, _) := xv0 in x7)
+ (ZRange.normalize args3) &&
+ is_bounded_by_bool
+ (let (x7, _) := xv1 in x7)
+ (ZRange.normalize args1)
+ then
+ Some
+ (#(Z_cast args)%expr @
+ v (Compile.reflect x6),
+ #(Z_cast r[0 ~> 0])%expr @
+ (##0)%expr)%expr_pat
+ else None);
+ Some (Base x7));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | @expr.App _ _ _ s2 _ ($_)%expr _ | @expr.App _ _ _ s2
+ _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s2 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s2 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
- end
- | @expr.App _ _ _ s _ (@expr.Ident _ _ _ t1 idc1) x3 =>
+ end;;
match x2 with
- | @expr.Ident _ _ _ t2 idc2 =>
- args <- invert_bind_args idc2 Raw.ident.Literal;
- args0 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc0 Raw.ident.Literal;
- args2 <- invert_bind_args idc Raw.ident.Literal;
+ | (@expr.Ident _ _ _ t4 idc4 @ @expr.Ident _ _ _ t5 idc5)%expr_pat =>
+ args <- invert_bind_args idc5 Raw.ident.Literal;
+ args0 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc2 Raw.ident.Literal;
+ args3 <- invert_bind_args idc1 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc0 Raw.ident.Literal;
+ args5 <- invert_bind_args idc Raw.ident.Z_cast;
match
pattern.type.unify_extracted
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) -> s) ->
+ ((((projT1 args4) -> (projT1 args2)) -> s1) ->
(projT1 args))%ptype
with
| Some (_, _, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) -> s) ->
+ ((((projT1 args4) -> (projT1 args2)) -> s1) ->
(projT1 args))%ptype
then
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
+ ##(projT2 args4);
xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s ℤ;
+ ##(projT2 args2);
+ v <- type.try_make_transport_cps s1 ℤ;
xv1 <- ident.unify pattern.ident.Literal
##(projT2 args);
- fv <- (x4 <- (if
- ((let (x4, _) := xv0 in x4) =? 0) &&
- ((let (x4, _) := xv1 in x4) =? 0) &&
- (ZRange.normalize args0 <=?
- r[0 ~> (let (x4, _) := xv in x4) -
- 1])%zrange
+ fv <- (x7 <- (if
+ ((let (x7, _) := xv0 in x7) =? 0) &&
+ ((let (x7, _) := xv1 in x7) =? 0) &&
+ (ZRange.normalize args1 <=?
+ r[0 ~> (let (x7, _) := xv in x7) -
+ 1])%zrange &&
+ is_bounded_by_bool
+ (let (x7, _) := xv in x7)
+ (ZRange.normalize args5) &&
+ is_bounded_by_bool
+ (let (x7, _) := xv0 in x7)
+ (ZRange.normalize args3) &&
+ is_bounded_by_bool
+ (let (x7, _) := xv1 in x7)
+ (ZRange.normalize args0)
then
Some
- (#(Z_cast args0)%expr @
- v (Compile.reflect x3),
+ (#(Z_cast args1)%expr @
+ v (Compile.reflect x5),
+ #(Z_cast r[0 ~> 0])%expr @
(##0)%expr)%expr_pat
else None);
- Some (Base x4));
+ Some (Base x7));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
+ | (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end
- | @expr.App _ _ _ s _ ($_)%expr _ | @expr.App _ _ _ s _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s _
+ | @expr.App _ _ _ s1 _ ($_)%expr _ | @expr.App _ _ _ s1 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s1 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s1 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
+ | (@expr.Ident _ _ _ t1 idc1 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end
+ | (@expr.Ident _ _ _ t idc @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t idc @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t idc @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat => None
| _ => None
end;;;
Base (#(Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat)%option
@@ -882,7 +1157,9 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
if type.type_beq base.type base.type.type_beq ℤ ℤ
then
fv <- (x0 <- (if lower range =? upper range
- then Some (##(lower range))%expr
+ then
+ Some
+ (#(Z_cast range)%expr @ (##(lower range))%expr)%expr_pat
else None);
Some (Base x0));
Some (fv0 <-- fv;
@@ -891,24 +1168,6 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
| None => None
end;;
match x with
- | @expr.Ident _ _ _ t idc =>
- args <- invert_bind_args idc Raw.ident.Literal;
- match pattern.type.unify_extracted ℤ (projT1 args) with
- | Some _ =>
- if type.type_beq base.type base.type.type_beq ℤ (projT1 args)
- then
- xv <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (x0 <- (if
- is_bounded_by_bool (let (x0, _) := xv in x0)
- range
- then Some (##(let (x0, _) := xv in x0))%expr
- else None);
- Some (Base x0));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
| @expr.App _ _ _ s _ (@expr.Ident _ _ _ t idc) x0 =>
args <- invert_bind_args idc Raw.ident.Z_cast;
match pattern.type.unify_extracted ℤ s with
@@ -931,126 +1190,6 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
| None => None
end
| @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.App _ _ _ s1 _ (@expr.Ident _ _ _ t1 idc1) x2) =>
- args <- invert_bind_args idc1 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_add;
- match
- pattern.type.unify_extracted (ℤ -> ℤ)%ptype
- ((projT1 args0) -> s1)%ptype
- with
- | Some (_, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
- ((projT1 args0) -> s1)%ptype
- then
- xv <- ident.unify pattern.ident.Literal ##(projT2 args0);
- v <- type.try_make_transport_cps s1 ℤ;
- fv <- (x3 <- (if
- ((let (x3, _) := xv in x3) =? 0) &&
- (ZRange.normalize args <=?
- ZRange.normalize range)%zrange
- then
- Some
- (#(Z_cast args)%expr @
- v (Compile.reflect x2))%expr_pat
- else None);
- Some (Base x3));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0)) (@expr.App _ _ _ s1 _ ($_)%expr _) |
- @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _) | @expr.App _ _ _ s
- _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.App _ _ _ s1 _ (_ @ _)%expr_pat _) | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.App _ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _) => None
- | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0)) #(_)%expr_pat | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0)) ($_)%expr | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0)) (@expr.Abs _ _ _ _ _ _) | @expr.App _
- _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0)) (@expr.LetIn _ _ _ _ _ _ _) => None
- | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.App _ _ _ s1 _ (@expr.Ident _ _ _ t0 idc0) x2))
- (@expr.Ident _ _ _ t1 idc1) =>
- args <- invert_bind_args idc1 Raw.ident.Literal;
- args0 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_add;
- match
- pattern.type.unify_extracted (ℤ -> ℤ)%ptype
- (s1 -> (projT1 args))%ptype
- with
- | Some (_, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq (ℤ -> ℤ)%ptype
- (s1 -> (projT1 args))%ptype
- then
- v <- type.try_make_transport_cps s1 ℤ;
- xv <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (x3 <- (if
- ((let (x3, _) := xv in x3) =? 0) &&
- (ZRange.normalize args0 <=?
- ZRange.normalize range)%zrange
- then
- Some
- (#(Z_cast args0)%expr @
- v (Compile.reflect x2))%expr_pat
- else None);
- Some (Base x3));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.App _ _ _ s1 _ (@expr.Ident _ _ _ t0 idc0) x2)) ($_)%expr |
- @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.App _ _ _ s1 _ (@expr.Ident _ _ _ t0 idc0) x2))
- (@expr.Abs _ _ _ _ _ _) | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.App _ _ _ s1 _ (@expr.Ident _ _ _ t0 idc0) x2))
- (_ @ _)%expr_pat | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.App _ _ _ s1 _ (@expr.Ident _ _ _ t0 idc0) x2))
- (@expr.LetIn _ _ _ _ _ _ _) => None
- | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.App _ _ _ s1 _ ($_)%expr _)) _ | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.App _ _ _ s1 _ (@expr.Abs _ _ _ _ _ _) _)) _ | @expr.App _ _
- _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.App _ _ _ s1 _ (_ @ _)%expr_pat _)) _ | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.App _ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _)) _ => None
- | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc) ($_)%expr) _ |
- @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Abs _ _ _ _ _ _)) _ | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.LetIn _ _ _ _ _ _ _)) _ => None
- | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _ (@expr.Ident _ _ _ t idc) x2) x1) x0 =>
_ <- invert_bind_args idc Raw.ident.Z_add_with_carry;
@@ -1086,7 +1225,8 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
_ | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _ (@expr.LetIn _ _ _ _ _ _ _) _) _) _ => None
- | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
+ | @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ #(_)%expr_pat _) _ |
+ @expr.App _ _ _ s _ (@expr.App _ _ _ s0 _ ($_)%expr _) _ | @expr.App
_ _ _ s _ (@expr.App _ _ _ s0 _ (@expr.Abs _ _ _ _ _ _) _) _ |
@expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.LetIn _ _ _ _ _ _ _) _) _ => None
@@ -1305,8 +1445,9 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
| _ => None
end;;
match x1 with
- | @expr.Ident _ _ _ t0 idc0 =>
- args <- invert_bind_args idc0 Raw.ident.Literal;
+ | (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat =>
+ args <- invert_bind_args idc1 Raw.ident.Literal;
+ args0 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_add_get_carry;
match
pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
@@ -1320,7 +1461,10 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
v <- type.try_make_transport_cps s1 ℤ;
xv <- ident.unify pattern.ident.Literal ##(projT2 args);
v0 <- type.try_make_transport_cps s ℤ;
- fv <- (if (let (x3, _) := xv in x3) <? 0
+ fv <- (if
+ ((let (x4, _) := xv in x4) <? 0) &&
+ is_bounded_by_bool (let (x4, _) := xv in x4)
+ (ZRange.normalize args0)
then
Some
(UnderLet
@@ -1330,7 +1474,8 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(#(Z_sub_get_borrow)%expr @
v (Compile.reflect x2) @
v0 (Compile.reflect x0) @
- (##(- (let (x3, _) := xv in x3))%Z)%expr))%expr_pat
+ (#(Z_cast (- args0))%expr @
+ (##(- (let (x4, _) := xv in x4))%Z)%expr)))%expr_pat
(fun vc : var (ℤ * ℤ)%etype =>
Base
(#(Z_cast (Datatypes.fst range))%expr @
@@ -1352,11 +1497,17 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
else None
| None => None
end
+ | (@expr.Ident _ _ _ t0 idc0 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end;;
match x0 with
- | @expr.Ident _ _ _ t0 idc0 =>
- args <- invert_bind_args idc0 Raw.ident.Literal;
+ | (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat =>
+ args <- invert_bind_args idc1 Raw.ident.Literal;
+ args0 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_add_get_carry;
match
pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
@@ -1371,7 +1522,10 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
v <- type.try_make_transport_cps s1 ℤ;
v0 <- type.try_make_transport_cps s0 ℤ;
xv <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (if (let (x3, _) := xv in x3) <? 0
+ fv <- (if
+ ((let (x4, _) := xv in x4) <? 0) &&
+ is_bounded_by_bool (let (x4, _) := xv in x4)
+ (ZRange.normalize args0)
then
Some
(UnderLet
@@ -1381,7 +1535,8 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(#(Z_sub_get_borrow)%expr @
v (Compile.reflect x2) @
v0 (Compile.reflect x1) @
- (##(- (let (x3, _) := xv in x3))%Z)%expr))%expr_pat
+ (#(Z_cast (- args0))%expr @
+ (##(- (let (x4, _) := xv in x4))%Z)%expr)))%expr_pat
(fun vc : var (ℤ * ℤ)%etype =>
Base
(#(Z_cast (Datatypes.fst range))%expr @
@@ -1403,6 +1558,11 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
else None
| None => None
end
+ | (@expr.Ident _ _ _ t0 idc0 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end;;
_ <- invert_bind_args idc Raw.ident.Z_add_get_carry;
@@ -1491,472 +1651,541 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(@expr.App _ _ _ s1 _
(@expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t idc) x3) x2) x1) x0 =>
(match x2 with
- | @expr.Ident _ _ _ t0 idc0 =>
- match x1 with
- | (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s5 _
- (@expr.Ident _ _ _ t3 idc3) x6))%expr_pat =>
- (args <- invert_bind_args idc3 Raw.ident.Z_cast;
- _ <- invert_bind_args idc2 Raw.ident.Z_opp;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args2)) -> s5) -> s)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args2)) -> s5) -> s)%ptype
- then
- v <- type.try_make_transport_cps s2 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v0 <- type.try_make_transport_cps s5 ℤ;
- v1 <- type.try_make_transport_cps s ℤ;
- fv <- (if
- ((let (x7, _) := xv in x7) =? 0) &&
- (ZRange.normalize args <=?
- - ZRange.normalize args1)%zrange
- then
- Some
- (UnderLet
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- (#(Z_sub_get_borrow)%expr @
- v (Compile.reflect x3) @
- v1 (Compile.reflect x0) @
- (#(Z_cast args)%expr @
- v0 (Compile.reflect x6))))%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(Z_cast (Datatypes.fst range))%expr @
- (#(fst)%expr @
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- ($vc)%expr)),
- #(Z_cast (Datatypes.snd range))%expr @
- (-
- (#(Z_cast (- Datatypes.snd range))%expr @
- (#(snd)%expr @
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- $vc)))%expr_pat)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end);;
- args <- invert_bind_args idc3 Raw.ident.Z_cast;
- _ <- invert_bind_args idc2 Raw.ident.Z_opp;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args2)) -> s5) -> s)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args2)) -> s5) -> s)%ptype
- then
- v <- type.try_make_transport_cps s2 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v0 <- type.try_make_transport_cps s5 ℤ;
- v1 <- type.try_make_transport_cps s ℤ;
- fv <- (if
- ((let (x7, _) := xv in x7) <? 0) &&
- (ZRange.normalize args <=?
- - ZRange.normalize args1)%zrange
- then
- Some
- (UnderLet
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- (#(Z_sub_with_get_borrow)%expr @
- v (Compile.reflect x3) @
- (##(- (let (x7, _) := xv in x7))%Z)%expr @
- v1 (Compile.reflect x0) @
- (#(Z_cast args)%expr @
- v0 (Compile.reflect x6))))%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(Z_cast (Datatypes.fst range))%expr @
- (#(fst)%expr @
+ | @expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t0 idc0) x4 =>
+ match x4 with
+ | @expr.Ident _ _ _ t1 idc1 =>
+ match x1 with
+ | (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (@expr.Ident _ _ _ t4 idc4) x7))%expr_pat =>
+ (args <- invert_bind_args idc4 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc3 Raw.ident.Z_opp;
+ args1 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc1 Raw.ident.Literal;
+ args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc
+ Raw.ident.Z_add_with_get_carry;
+ match
+ pattern.type.unify_extracted
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((s2 -> (projT1 args2)) -> s6) -> s)%ptype
+ with
+ | Some (_, _, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((s2 -> (projT1 args2)) -> s6) -> s)%ptype
+ then
+ v <- type.try_make_transport_cps s2 ℤ;
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args2);
+ v0 <- type.try_make_transport_cps s6 ℤ;
+ v1 <- type.try_make_transport_cps s ℤ;
+ fv <- (if
+ ((let (x8, _) := xv in x8) =? 0) &&
+ (ZRange.normalize args <=?
+ - ZRange.normalize args1)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args3)
+ then
+ Some
+ (UnderLet
(#(Z_cast2
(Datatypes.fst range,
- Datatypes.snd range))%expr @
- ($vc)%expr)),
- #(Z_cast (Datatypes.snd range))%expr @
- (-
- (#(Z_cast (- Datatypes.snd range))%expr @
- (#(snd)%expr @
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- $vc)))%expr_pat)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s5 _ ($_)%expr
- _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s5 _
- (@expr.Abs _ _ _ _ _ _) _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s5 _ (_ @ _) _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s5 _
- (@expr.LetIn _ _ _ _ _ _ _) _))%expr_pat => None
- | (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ #(_)))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.Abs _ _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t1 idc1 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ (_ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
- None
- | _ => None
- end;;
- match x0 with
- | (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s5 _
- (@expr.Ident _ _ _ t3 idc3) x6))%expr_pat =>
- (args <- invert_bind_args idc3 Raw.ident.Z_cast;
- _ <- invert_bind_args idc2 Raw.ident.Z_opp;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args2)) -> s0) -> s5)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
+ (#(Z_sub_get_borrow)%expr @
+ v (Compile.reflect x3) @
+ v1 (Compile.reflect x0) @
+ (#(Z_cast args)%expr @
+ v0 (Compile.reflect x7))))%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(Z_cast (Datatypes.fst range))%expr @
+ (#(fst)%expr @
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ ($vc)%expr)),
+ #(Z_cast (Datatypes.snd range))%expr @
+ (-
+ (#(Z_cast
+ (- Datatypes.snd range))%expr @
+ (#(snd)%expr @
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ $vc)))%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end);;
+ args <- invert_bind_args idc4 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc3 Raw.ident.Z_opp;
+ args1 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc1 Raw.ident.Literal;
+ args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc
+ Raw.ident.Z_add_with_get_carry;
+ match
+ pattern.type.unify_extracted
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args2)) -> s0) -> s5)%ptype
- then
- v <- type.try_make_transport_cps s2 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v0 <- type.try_make_transport_cps s0 ℤ;
- v1 <- type.try_make_transport_cps s5 ℤ;
- fv <- (if
- ((let (x7, _) := xv in x7) =? 0) &&
- (ZRange.normalize args <=?
- - ZRange.normalize args1)%zrange
- then
- Some
- (UnderLet
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- (#(Z_sub_get_borrow)%expr @
- v (Compile.reflect x3) @
- v0 (Compile.reflect x1) @
- (#(Z_cast args)%expr @
- v1 (Compile.reflect x6))))%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(Z_cast (Datatypes.fst range))%expr @
- (#(fst)%expr @
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- ($vc)%expr)),
- #(Z_cast (Datatypes.snd range))%expr @
- (-
- (#(Z_cast (- Datatypes.snd range))%expr @
- (#(snd)%expr @
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- $vc)))%expr_pat)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end);;
- args <- invert_bind_args idc3 Raw.ident.Z_cast;
- _ <- invert_bind_args idc2 Raw.ident.Z_opp;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args2)) -> s0) -> s5)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args2)) -> s0) -> s5)%ptype
- then
- v <- type.try_make_transport_cps s2 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v0 <- type.try_make_transport_cps s0 ℤ;
- v1 <- type.try_make_transport_cps s5 ℤ;
- fv <- (if
- ((let (x7, _) := xv in x7) <? 0) &&
- (ZRange.normalize args <=?
- - ZRange.normalize args1)%zrange
- then
- Some
- (UnderLet
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- (#(Z_sub_with_get_borrow)%expr @
- v (Compile.reflect x3) @
- (##(- (let (x7, _) := xv in x7))%Z)%expr @
- v0 (Compile.reflect x1) @
- (#(Z_cast args)%expr @
- v1 (Compile.reflect x6))))%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(Z_cast (Datatypes.fst range))%expr @
- (#(fst)%expr @
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- ($vc)%expr)),
- #(Z_cast (Datatypes.snd range))%expr @
- (-
- (#(Z_cast (- Datatypes.snd range))%expr @
- (#(snd)%expr @
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- $vc)))%expr_pat)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s5 _ ($_)%expr
- _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s5 _
- (@expr.Abs _ _ _ _ _ _) _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s5 _ (_ @ _) _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.App _ _ _ s5 _
- (@expr.LetIn _ _ _ _ _ _ _) _))%expr_pat => None
- | (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ #(_)))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.Abs _ _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @
- (@expr.Ident _ _ _ t2 idc2 @ @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t1 idc1 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ (_ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t1 idc1 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
- None
- | _ => None
- end;;
- match x1 with
- | @expr.Ident _ _ _ t1 idc1 =>
- args <- invert_bind_args idc1 Raw.ident.Literal;
- args0 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args0)) -> (projT1 args)) -> s)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args0)) -> (projT1 args)) -> s)%ptype
- then
- v <- type.try_make_transport_cps s2 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v0 <- type.try_make_transport_cps s ℤ;
- fv <- (if
- ((let (x4, _) := xv0 in x4) <=? 0) &&
- ((let (x4, _) := xv in x4) <=? 0) &&
- ((let (x4, _) := xv0 in x4) +
- (let (x4, _) := xv in x4) <? 0)
- then
- Some
- (UnderLet
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- (#(Z_sub_with_get_borrow)%expr @
- v (Compile.reflect x3) @
- (##(- (let (x4, _) := xv in x4))%Z)%expr @
- v0 (Compile.reflect x0) @
- (##(- (let (x4, _) := xv0 in x4))%Z)%expr))%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(Z_cast (Datatypes.fst range))%expr @
- (#(fst)%expr @
+ (((s2 -> (projT1 args2)) -> s6) -> s)%ptype
+ with
+ | Some (_, _, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((s2 -> (projT1 args2)) -> s6) -> s)%ptype
+ then
+ v <- type.try_make_transport_cps s2 ℤ;
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args2);
+ v0 <- type.try_make_transport_cps s6 ℤ;
+ v1 <- type.try_make_transport_cps s ℤ;
+ fv <- (if
+ ((let (x8, _) := xv in x8) <? 0) &&
+ (ZRange.normalize args <=?
+ - ZRange.normalize args1)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args3)
+ then
+ Some
+ (UnderLet
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ (#(Z_sub_with_get_borrow)%expr @
+ v (Compile.reflect x3) @
+ (#(Z_cast (- args3))%expr @
+ (##(- (let (x8, _) := xv in x8))%Z)%expr) @
+ v1 (Compile.reflect x0) @
+ (#(Z_cast args)%expr @
+ v0 (Compile.reflect x7))))%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(Z_cast (Datatypes.fst range))%expr @
+ (#(fst)%expr @
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ ($vc)%expr)),
+ #(Z_cast (Datatypes.snd range))%expr @
+ (-
+ (#(Z_cast (- Datatypes.snd range))%expr @
+ (#(snd)%expr @
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ $vc)))%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ ($_)%expr _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (@expr.Abs _ _ _ _ _ _) _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (_ @ _) _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (@expr.LetIn _ _ _ _ _ _ _) _))%expr_pat => None
+ | (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ #(_)))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.Abs _ _ _ _ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
+ None
+ | (@expr.Ident _ _ _ t2 idc2 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ (_ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
+ | _ => None
+ end;;
+ match x0 with
+ | (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (@expr.Ident _ _ _ t4 idc4) x7))%expr_pat =>
+ (args <- invert_bind_args idc4 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc3 Raw.ident.Z_opp;
+ args1 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc1 Raw.ident.Literal;
+ args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc
+ Raw.ident.Z_add_with_get_carry;
+ match
+ pattern.type.unify_extracted
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((s2 -> (projT1 args2)) -> s0) -> s6)%ptype
+ with
+ | Some (_, _, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((s2 -> (projT1 args2)) -> s0) -> s6)%ptype
+ then
+ v <- type.try_make_transport_cps s2 ℤ;
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args2);
+ v0 <- type.try_make_transport_cps s0 ℤ;
+ v1 <- type.try_make_transport_cps s6 ℤ;
+ fv <- (if
+ ((let (x8, _) := xv in x8) =? 0) &&
+ (ZRange.normalize args <=?
+ - ZRange.normalize args1)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args3)
+ then
+ Some
+ (UnderLet
(#(Z_cast2
(Datatypes.fst range,
- Datatypes.snd range))%expr @
- ($vc)%expr)),
- #(Z_cast (Datatypes.snd range))%expr @
- (-
- (#(Z_cast (- Datatypes.snd range))%expr @
- (#(snd)%expr @
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- $vc)))%expr_pat)%expr)%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end;;
- match x0 with
- | @expr.Ident _ _ _ t1 idc1 =>
+ (#(Z_sub_get_borrow)%expr @
+ v (Compile.reflect x3) @
+ v0 (Compile.reflect x1) @
+ (#(Z_cast args)%expr @
+ v1 (Compile.reflect x7))))%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(Z_cast (Datatypes.fst range))%expr @
+ (#(fst)%expr @
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ ($vc)%expr)),
+ #(Z_cast (Datatypes.snd range))%expr @
+ (-
+ (#(Z_cast
+ (- Datatypes.snd range))%expr @
+ (#(snd)%expr @
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ $vc)))%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end);;
+ args <- invert_bind_args idc4 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc3 Raw.ident.Z_opp;
+ args1 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc1 Raw.ident.Literal;
+ args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc
+ Raw.ident.Z_add_with_get_carry;
+ match
+ pattern.type.unify_extracted
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((s2 -> (projT1 args2)) -> s0) -> s6)%ptype
+ with
+ | Some (_, _, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((s2 -> (projT1 args2)) -> s0) -> s6)%ptype
+ then
+ v <- type.try_make_transport_cps s2 ℤ;
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args2);
+ v0 <- type.try_make_transport_cps s0 ℤ;
+ v1 <- type.try_make_transport_cps s6 ℤ;
+ fv <- (if
+ ((let (x8, _) := xv in x8) <? 0) &&
+ (ZRange.normalize args <=?
+ - ZRange.normalize args1)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args3)
+ then
+ Some
+ (UnderLet
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ (#(Z_sub_with_get_borrow)%expr @
+ v (Compile.reflect x3) @
+ (#(Z_cast (- args3))%expr @
+ (##(- (let (x8, _) := xv in x8))%Z)%expr) @
+ v0 (Compile.reflect x1) @
+ (#(Z_cast args)%expr @
+ v1 (Compile.reflect x7))))%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(Z_cast (Datatypes.fst range))%expr @
+ (#(fst)%expr @
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ ($vc)%expr)),
+ #(Z_cast (Datatypes.snd range))%expr @
+ (-
+ (#(Z_cast (- Datatypes.snd range))%expr @
+ (#(snd)%expr @
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ $vc)))%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ ($_)%expr _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (@expr.Abs _ _ _ _ _ _) _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (_ @ _) _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (@expr.LetIn _ _ _ _ _ _ _) _))%expr_pat => None
+ | (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ #(_)))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.Abs _ _ _ _ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
+ None
+ | (@expr.Ident _ _ _ t2 idc2 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ (_ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
+ | _ => None
+ end;;
+ match x1 with
+ | (@expr.Ident _ _ _ t2 idc2 @ @expr.Ident _ _ _ t3 idc3)%expr_pat =>
+ args <- invert_bind_args idc3 Raw.ident.Literal;
+ args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc1 Raw.ident.Literal;
+ args2 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc
+ Raw.ident.Z_add_with_get_carry;
+ match
+ pattern.type.unify_extracted
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((s2 -> (projT1 args1)) -> (projT1 args)) -> s)%ptype
+ with
+ | Some (_, _, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((s2 -> (projT1 args1)) -> (projT1 args)) -> s)%ptype
+ then
+ v <- type.try_make_transport_cps s2 ℤ;
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v0 <- type.try_make_transport_cps s ℤ;
+ fv <- (if
+ ((let (x6, _) := xv0 in x6) <=? 0) &&
+ ((let (x6, _) := xv in x6) <=? 0) &&
+ ((let (x6, _) := xv0 in x6) +
+ (let (x6, _) := xv in x6) <? 0) &&
+ is_bounded_by_bool
+ (let (x6, _) := xv0 in x6)
+ (ZRange.normalize args0) &&
+ is_bounded_by_bool
+ (let (x6, _) := xv in x6)
+ (ZRange.normalize args2)
+ then
+ Some
+ (UnderLet
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ (#(Z_sub_with_get_borrow)%expr @
+ v (Compile.reflect x3) @
+ (#(Z_cast (- args2))%expr @
+ (##(- (let (x6, _) := xv in x6))%Z)%expr) @
+ v0 (Compile.reflect x0) @
+ (#(Z_cast (- args0))%expr @
+ (##(- (let (x6, _) := xv0 in x6))%Z)%expr)))%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(Z_cast (Datatypes.fst range))%expr @
+ (#(fst)%expr @
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ ($vc)%expr)),
+ #(Z_cast (Datatypes.snd range))%expr @
+ (-
+ (#(Z_cast (- Datatypes.snd range))%expr @
+ (#(snd)%expr @
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ $vc)))%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | (@expr.Ident _ _ _ t2 idc2 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
+ | _ => None
+ end;;
+ match x0 with
+ | (@expr.Ident _ _ _ t2 idc2 @ @expr.Ident _ _ _ t3 idc3)%expr_pat =>
+ args <- invert_bind_args idc3 Raw.ident.Literal;
+ args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc1 Raw.ident.Literal;
+ args2 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc
+ Raw.ident.Z_add_with_get_carry;
+ match
+ pattern.type.unify_extracted
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((s2 -> (projT1 args1)) -> s0) -> (projT1 args))%ptype
+ with
+ | Some (_, _, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ (((s2 -> (projT1 args1)) -> s0) -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s2 ℤ;
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v0 <- type.try_make_transport_cps s0 ℤ;
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (if
+ ((let (x6, _) := xv0 in x6) <=? 0) &&
+ ((let (x6, _) := xv in x6) <=? 0) &&
+ ((let (x6, _) := xv0 in x6) +
+ (let (x6, _) := xv in x6) <? 0) &&
+ is_bounded_by_bool
+ (let (x6, _) := xv0 in x6)
+ (ZRange.normalize args0) &&
+ is_bounded_by_bool
+ (let (x6, _) := xv in x6)
+ (ZRange.normalize args2)
+ then
+ Some
+ (UnderLet
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ (#(Z_sub_with_get_borrow)%expr @
+ v (Compile.reflect x3) @
+ (#(Z_cast (- args2))%expr @
+ (##(- (let (x6, _) := xv in x6))%Z)%expr) @
+ v0 (Compile.reflect x1) @
+ (#(Z_cast (- args0))%expr @
+ (##(- (let (x6, _) := xv0 in x6))%Z)%expr)))%expr_pat
+ (fun vc : var (ℤ * ℤ)%etype =>
+ Base
+ (#(Z_cast (Datatypes.fst range))%expr @
+ (#(fst)%expr @
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ ($vc)%expr)),
+ #(Z_cast (Datatypes.snd range))%expr @
+ (-
+ (#(Z_cast (- Datatypes.snd range))%expr @
+ (#(snd)%expr @
+ (#(Z_cast2
+ (Datatypes.fst range,
+ - Datatypes.snd range))%expr @
+ $vc)))%expr_pat)%expr)%expr_pat))
+ else None);
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | (@expr.Ident _ _ _ t2 idc2 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
+ | _ => None
+ end;;
args <- invert_bind_args idc1 Raw.ident.Literal;
- args0 <- invert_bind_args idc0 Raw.ident.Literal;
+ args0 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_add_with_get_carry;
match
pattern.type.unify_extracted (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args0)) -> s0) -> (projT1 args))%ptype
+ (((s2 -> (projT1 args)) -> s0) -> s)%ptype
with
| Some (_, _, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args0)) -> s0) -> (projT1 args))%ptype
+ (((s2 -> (projT1 args)) -> s0) -> s)%ptype
then
v <- type.try_make_transport_cps s2 ℤ;
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
+ ##(projT2 args);
v0 <- type.try_make_transport_cps s0 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
+ v1 <- type.try_make_transport_cps s ℤ;
fv <- (if
- ((let (x4, _) := xv0 in x4) <=? 0) &&
- ((let (x4, _) := xv in x4) <=? 0) &&
- ((let (x4, _) := xv0 in x4) +
- (let (x4, _) := xv in x4) <? 0)
+ ((let (x5, _) := xv in x5) =? 0) &&
+ is_bounded_by_bool (let (x5, _) := xv in x5)
+ (ZRange.normalize args0)
then
Some
(UnderLet
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- (#(Z_sub_with_get_borrow)%expr @
+ (#(Z_cast2 range)%expr @
+ (#(Z_add_get_carry)%expr @
v (Compile.reflect x3) @
- (##(- (let (x4, _) := xv in x4))%Z)%expr @
v0 (Compile.reflect x1) @
- (##(- (let (x4, _) := xv0 in x4))%Z)%expr))%expr_pat
+ v1 (Compile.reflect x0)))%expr_pat
(fun vc : var (ℤ * ℤ)%etype =>
Base
(#(Z_cast (Datatypes.fst range))%expr @
(#(fst)%expr @
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- ($vc)%expr)),
+ (#(Z_cast2 range)%expr @ ($vc)%expr)),
#(Z_cast (Datatypes.snd range))%expr @
- (-
- (#(Z_cast (- Datatypes.snd range))%expr @
- (#(snd)%expr @
- (#(Z_cast2
- (Datatypes.fst range,
- - Datatypes.snd range))%expr @
- $vc)))%expr_pat)%expr)%expr_pat))
+ (#(snd)%expr @
+ (#(Z_cast2 range)%expr @ ($vc)%expr)))%expr_pat))
else None);
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | _ => None
- end;;
- args <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args)) -> s0) -> s)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s2 -> (projT1 args)) -> s0) -> s)%ptype
- then
- v <- type.try_make_transport_cps s2 ℤ;
- xv <- ident.unify pattern.ident.Literal ##(projT2 args);
- v0 <- type.try_make_transport_cps s0 ℤ;
- v1 <- type.try_make_transport_cps s ℤ;
- fv <- (if (let (x4, _) := xv in x4) =? 0
- then
- Some
- (UnderLet
- (#(Z_cast2 range)%expr @
- (#(Z_add_get_carry)%expr @
- v (Compile.reflect x3) @
- v0 (Compile.reflect x1) @
- v1 (Compile.reflect x0)))%expr_pat
- (fun vc : var (ℤ * ℤ)%etype =>
- Base
- (#(Z_cast (Datatypes.fst range))%expr @
- (#(fst)%expr @
- (#(Z_cast2 range)%expr @ ($vc)%expr)),
- #(Z_cast (Datatypes.snd range))%expr @
- (#(snd)%expr @
- (#(Z_cast2 range)%expr @ ($vc)%expr)))%expr_pat))
- else None);
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t0 idc0) x4 =>
- match x4 with
| (@expr.Ident _ _ _ t1 idc1 @ @expr.App _ _ _ s5 _
(@expr.Ident _ _ _ t2 idc2) x6)%expr_pat =>
match x1 with
@@ -2160,11 +2389,12 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
| _ => None
end;;
match x1 with
- | @expr.Ident _ _ _ t3 idc3 =>
- args <- invert_bind_args idc3 Raw.ident.Literal;
- args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ | (@expr.Ident _ _ _ t3 idc3 @ @expr.Ident _ _ _ t4 idc4)%expr_pat =>
+ args <- invert_bind_args idc4 Raw.ident.Literal;
+ args0 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc2 Raw.ident.Z_cast;
_ <- invert_bind_args idc1 Raw.ident.Z_opp;
- args2 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc
Raw.ident.Z_add_with_get_carry;
match
@@ -2184,9 +2414,12 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
##(projT2 args);
v1 <- type.try_make_transport_cps s ℤ;
fv <- (if
- ((let (x7, _) := xv in x7) <=? 0) &&
- (ZRange.normalize args0 <=?
- - ZRange.normalize args2)%zrange
+ ((let (x8, _) := xv in x8) <=? 0) &&
+ (ZRange.normalize args1 <=?
+ - ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args0)
then
Some
(UnderLet
@@ -2195,10 +2428,11 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
- Datatypes.snd range))%expr @
(#(Z_sub_with_get_borrow)%expr @
v (Compile.reflect x3) @
- (#(Z_cast args0)%expr @
+ (#(Z_cast args1)%expr @
v0 (Compile.reflect x6)) @
v1 (Compile.reflect x0) @
- (##(- (let (x7, _) := xv in x7))%Z)%expr))%expr_pat
+ (#(Z_cast (- args0))%expr @
+ (##(- (let (x8, _) := xv in x8))%Z)%expr)))%expr_pat
(fun vc : var (ℤ * ℤ)%etype =>
Base
(#(Z_cast (Datatypes.fst range))%expr @
@@ -2221,14 +2455,20 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
else None
| None => None
end
+ | (@expr.Ident _ _ _ t3 idc3 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end;;
match x0 with
- | @expr.Ident _ _ _ t3 idc3 =>
- args <- invert_bind_args idc3 Raw.ident.Literal;
- args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ | (@expr.Ident _ _ _ t3 idc3 @ @expr.Ident _ _ _ t4 idc4)%expr_pat =>
+ args <- invert_bind_args idc4 Raw.ident.Literal;
+ args0 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc2 Raw.ident.Z_cast;
_ <- invert_bind_args idc1 Raw.ident.Z_opp;
- args2 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc
Raw.ident.Z_add_with_get_carry;
match
@@ -2248,9 +2488,12 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
xv <- ident.unify pattern.ident.Literal
##(projT2 args);
fv <- (if
- ((let (x7, _) := xv in x7) <=? 0) &&
- (ZRange.normalize args0 <=?
- - ZRange.normalize args2)%zrange
+ ((let (x8, _) := xv in x8) <=? 0) &&
+ (ZRange.normalize args1 <=?
+ - ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args0)
then
Some
(UnderLet
@@ -2259,10 +2502,11 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
- Datatypes.snd range))%expr @
(#(Z_sub_with_get_borrow)%expr @
v (Compile.reflect x3) @
- (#(Z_cast args0)%expr @
+ (#(Z_cast args1)%expr @
v0 (Compile.reflect x6)) @
v1 (Compile.reflect x1) @
- (##(- (let (x7, _) := xv in x7))%Z)%expr))%expr_pat
+ (#(Z_cast (- args0))%expr @
+ (##(- (let (x8, _) := xv in x8))%Z)%expr)))%expr_pat
(fun vc : var (ℤ * ℤ)%etype =>
Base
(#(Z_cast (Datatypes.fst range))%expr @
@@ -2285,6 +2529,11 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
else None
| None => None
end
+ | (@expr.Ident _ _ _ t3 idc3 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end
| (@expr.Ident _ _ _ t1 idc1 @ @expr.App _ _ _ s5 _ ($_)%expr _)%expr_pat |
@@ -2301,71 +2550,104 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
| _ => None
end;;
match x3 with
- | @expr.Ident _ _ _ t1 idc1 =>
+ | (@expr.Ident _ _ _ t1 idc1 @ @expr.Ident _ _ _ t2 idc2)%expr_pat =>
match x1 with
- | @expr.Ident _ _ _ t2 idc2 =>
+ | (@expr.Ident _ _ _ t3 idc3 @ @expr.Ident _ _ _ t4 idc4)%expr_pat =>
match x0 with
- | @expr.Ident _ _ _ t3 idc3 =>
- args <- invert_bind_args idc3 Raw.ident.Literal;
- args0 <- invert_bind_args idc2 Raw.ident.Literal;
- args1 <- invert_bind_args idc1 Raw.ident.Literal;
- args2 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.Ident _ _ _ t6
+ idc6)%expr_pat =>
+ args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Literal;
+ args2 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc2 Raw.ident.Literal;
+ args4 <- invert_bind_args idc1 Raw.ident.Z_cast;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc
Raw.ident.Z_add_with_get_carry;
match
pattern.type.unify_extracted
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args1) -> s3) -> (projT1 args0)) ->
+ ((((projT1 args3) -> s3) -> (projT1 args1)) ->
(projT1 args))%ptype
with
| Some (_, _, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args1) -> s3) -> (projT1 args0)) ->
+ ((((projT1 args3) -> s3) -> (projT1 args1)) ->
(projT1 args))%ptype
then
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
+ ##(projT2 args3);
v <- type.try_make_transport_cps s3 ℤ;
xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
+ ##(projT2 args1);
xv1 <- ident.unify pattern.ident.Literal
##(projT2 args);
fv <- (if
- ((let (x5, _) := xv0 in x5) =? 0) &&
- ((let (x5, _) := xv1 in x5) =? 0) &&
- (ZRange.normalize args2 <=?
- r[0 ~> (let (x5, _) := xv in x5) - 1])%zrange &&
+ ((let (x8, _) := xv0 in x8) =? 0) &&
+ ((let (x8, _) := xv1 in x8) =? 0) &&
+ (ZRange.normalize args5 <=?
+ r[0 ~> (let (x8, _) := xv in x8) - 1])%zrange &&
is_bounded_by_bool 0
- (Datatypes.snd range)
+ (Datatypes.snd range) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args4) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args2) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv1 in x8)
+ (ZRange.normalize args0)
then
Some
(UnderLet
(#(Z_cast2 range)%expr @
(#(Z_add_with_get_carry)%expr @
- (##(let (x5, _) := xv in x5))%expr @
- (#(Z_cast args2)%expr @
+ (#(Z_cast args4)%expr @
+ (##(let (x8, _) := xv in x8))%expr) @
+ (#(Z_cast args5)%expr @
v (Compile.reflect x4)) @
- (##(let (x5, _) := xv0 in x5))%expr @
- (##(let (x5, _) := xv1 in x5))%expr))%expr_pat
+ (#(Z_cast args2)%expr @
+ (##(let (x8, _) := xv0 in x8))%expr) @
+ (#(Z_cast args0)%expr @
+ (##(let (x8, _) := xv1 in x8))%expr)))%expr_pat
(fun vc : var (ℤ * ℤ)%etype =>
Base
(#(Z_cast
(Datatypes.fst range))%expr @
(#(fst)%expr @
(#(Z_cast2 range)%expr @
- ($vc)%expr)), (##0)%expr)%expr_pat))
+ ($vc)%expr)),
+ #(Z_cast r[0 ~> 0])%expr @
+ (##0)%expr)%expr_pat))
else None);
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
+ | (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _
+ _)%expr_pat => None
| _ => None
end
+ | (@expr.Ident _ _ _ t3 idc3 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end
+ | (@expr.Ident _ _ _ t1 idc1 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t1 idc1 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end
| @expr.App _ _ _ s3 _ ($_)%expr _ | @expr.App _ _ _ s3 _
diff --git a/src/fancy_with_casts_rewrite_head.out b/src/fancy_with_casts_rewrite_head.out
index 551fd7e18..c5f1388d8 100644
--- a/src/fancy_with_casts_rewrite_head.out
+++ b/src/fancy_with_casts_rewrite_head.out
@@ -178,1232 +178,1890 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((match x with
| @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.Ident _ _ _ t1 idc1 @ (@expr.Ident _ _ _ t2 idc2 @ x4 @ x3))%expr_pat =>
- match x4 with
- | @expr.Ident _ _ _ t3 idc3 =>
- match x3 with
- | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t4 idc4) x5 =>
- (args <- invert_bind_args idc4 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc3 Raw.ident.Literal;
- _ <- invert_bind_args idc2 Raw.ident.Z_land;
- args2 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted (ℤ -> ℤ -> ℤ)%ptype
- ((projT1 args3) -> (projT1 args0) -> s4)%ptype
- with
- | Some (_, (_, _))%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (ℤ -> ℤ -> ℤ)%ptype
- ((projT1 args3) -> (projT1 args0) -> s4)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args3);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s4 ℤ;
- fv <- (x6 <- (if
- ((let (x6, _) := xv0 in x6) =?
- 2
- ^ (2 *
- Z.log2_up
- (let (x6, _) := xv0 in x6) / 2) -
- 1) &&
- (ZRange.normalize
- (ZRange.constant
- (let (x6, _) := xv0 in x6)) &'
- ZRange.normalize args <=?
- ZRange.normalize args2)%zrange
- then
- x6 <- invert_low
- (2 *
- Z.log2_up
- (let (x6, _) := xv0 in x6))
- (let (x6, _) := xv in x6);
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let (x7, _) := xv0 in x7)))%expr @
- ((##x6)%expr,
- #(Z_cast args)%expr @
- v (Compile.reflect x5))))%expr_pat
- else None);
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end);;
- args <- invert_bind_args idc4 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc3 Raw.ident.Literal;
- _ <- invert_bind_args idc2 Raw.ident.Z_land;
- args2 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted (ℤ -> ℤ -> ℤ)%ptype
- ((projT1 args3) -> (projT1 args0) -> s4)%ptype
- with
- | Some (_, (_, _))%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (ℤ -> ℤ -> ℤ)%ptype
- ((projT1 args3) -> (projT1 args0) -> s4)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args3);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s4 ℤ;
- fv <- (x6 <- (if
- ((let (x6, _) := xv0 in x6) =?
- 2
- ^ (2 *
- Z.log2_up
- (let (x6, _) := xv0 in x6) / 2) -
- 1) &&
- (ZRange.normalize
- (ZRange.constant
- (let (x6, _) := xv0 in x6)) &'
- ZRange.normalize args <=?
- ZRange.normalize args2)%zrange
- then
- x6 <- invert_high
- (2 *
- Z.log2_up
- (let (x6, _) := xv0 in x6))
- (let (x6, _) := xv in x6);
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mulhl
- (2 *
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Ident _ _ _ t2 idc2 @ (@expr.Ident _ _ _ t3 idc3 @ x5 @ x4))%expr_pat =>
+ match x5 with
+ | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t4 idc4) x6 =>
+ match x4 with
+ | @expr.App _ _ _ s6 _ (@expr.Ident _ _ _ t5 idc5) x7 =>
+ match x7 with
+ | @expr.Ident _ _ _ t6 idc6 =>
+ args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc3 Raw.ident.Z_land;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc1 Raw.ident.Literal;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args4) -> s5 -> (projT1 args))%ptype
+ with
+ | Some (_, (_, _))%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args4) -> s5 -> (projT1 args))%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args4);
+ v <- type.try_make_transport_cps s5 ℤ;
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (x8 <- (if
+ ((let (x8, _) := xv0 in x8) =?
+ 2
+ ^ (2 *
+ Z.log2_up
+ (let (x8, _) := xv0 in x8) /
+ 2) - 1) &&
+ (ZRange.normalize args1 &'
+ ZRange.normalize
+ (ZRange.constant
+ (let (x8, _) := xv0 in x8)) <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args5) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args0)
+ then
+ x8 <- invert_low
+ (2 *
+ Z.log2_up
+ (let (x8, _) := xv0 in
+ x8))
+ (let (x8, _) := xv in x8);
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mulll
+ (2 *
+ Z.log2_up
+ (let (x9, _) := xv0 in
+ x9)))%expr @
+ ((##x8)%expr,
+ #(Z_cast args1)%expr @
+ v (Compile.reflect x6))))%expr_pat
+ else None);
+ Some (Base x8));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | _ => None
+ end;;
+ match x6 with
+ | @expr.Ident _ _ _ t6 idc6 =>
+ (args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc3 Raw.ident.Z_land;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc1 Raw.ident.Literal;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args4) -> (projT1 args) -> s6)%ptype
+ with
+ | Some (_, (_, _))%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args4) -> (projT1 args) -> s6)%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args4);
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v <- type.try_make_transport_cps s6 ℤ;
+ fv <- (x8 <- (if
+ ((let (x8, _) := xv0 in x8) =?
+ 2
+ ^ (2 *
Z.log2_up
- (let (x7, _) := xv0 in x7)))%expr @
- ((##x6)%expr,
- #(Z_cast args)%expr @
- v (Compile.reflect x5))))%expr_pat
- else None);
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
+ (let (x8, _) := xv0 in x8) /
+ 2) - 1) &&
+ (ZRange.normalize
+ (ZRange.constant
+ (let (x8, _) := xv0 in x8)) &'
+ ZRange.normalize args0 <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args5) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args1)
+ then
+ x8 <- invert_low
+ (2 *
+ Z.log2_up
+ (let (x8, _) := xv0 in
+ x8))
+ (let (x8, _) := xv in x8);
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mulll
+ (2 *
+ Z.log2_up
+ (let (x9, _) := xv0 in
+ x9)))%expr @
+ ((##x8)%expr,
+ #(Z_cast args0)%expr @
+ v (Compile.reflect x7))))%expr_pat
+ else None);
+ Some (Base x8));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end);;
+ args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc3 Raw.ident.Z_land;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc1 Raw.ident.Literal;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args4) -> (projT1 args) -> s6)%ptype
+ with
+ | Some (_, (_, _))%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args4) -> (projT1 args) -> s6)%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args4);
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ v <- type.try_make_transport_cps s6 ℤ;
+ fv <- (x8 <- (if
+ ((let (x8, _) := xv0 in x8) =?
+ 2
+ ^ (2 *
+ Z.log2_up
+ (let (x8, _) := xv0 in x8) /
+ 2) - 1) &&
+ (ZRange.normalize
+ (ZRange.constant
+ (let (x8, _) := xv0 in x8)) &'
+ ZRange.normalize args0 <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args5) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args1)
+ then
+ x8 <- invert_high
+ (2 *
+ Z.log2_up
+ (let (x8, _) := xv0 in
+ x8))
+ (let (x8, _) := xv in x8);
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mulhl
+ (2 *
+ Z.log2_up
+ (let (x9, _) := xv0 in
+ x9)))%expr @
+ ((##x8)%expr,
+ #(Z_cast args0)%expr @
+ v (Compile.reflect x7))))%expr_pat
+ else None);
+ Some (Base x8));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | _ => None
+ end;;
+ match x7 with
+ | @expr.Ident _ _ _ t6 idc6 =>
+ args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc3 Raw.ident.Z_land;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc1 Raw.ident.Literal;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args4) -> s5 -> (projT1 args))%ptype
+ with
+ | Some (_, (_, _))%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ (ℤ -> ℤ -> ℤ)%ptype
+ ((projT1 args4) -> s5 -> (projT1 args))%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args4);
+ v <- type.try_make_transport_cps s5 ℤ;
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (x8 <- (if
+ ((let (x8, _) := xv0 in x8) =?
+ 2
+ ^ (2 *
+ Z.log2_up
+ (let (x8, _) := xv0 in x8) /
+ 2) - 1) &&
+ (ZRange.normalize args1 &'
+ ZRange.normalize
+ (ZRange.constant
+ (let (x8, _) := xv0 in x8)) <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args5) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args0)
+ then
+ x8 <- invert_high
+ (2 *
+ Z.log2_up
+ (let (x8, _) := xv0 in
+ x8))
+ (let (x8, _) := xv in x8);
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mulhl
+ (2 *
+ Z.log2_up
+ (let (x9, _) := xv0 in
+ x9)))%expr @
+ ((##x8)%expr,
+ #(Z_cast args1)%expr @
+ v (Compile.reflect x6))))%expr_pat
+ else None);
+ Some (Base x8));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | _ => None
end
- | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s4 _
+ | @expr.App _ _ _ s6 _ ($_)%expr _ | @expr.App _ _ _ s6 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s6 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s6 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
- | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t3 idc3) x5 =>
- match x3 with
- | @expr.Ident _ _ _ t4 idc4 =>
- (args <- invert_bind_args idc4 Raw.ident.Literal;
- args0 <- invert_bind_args idc3 Raw.ident.Z_cast;
- _ <- invert_bind_args idc2 Raw.ident.Z_land;
- args2 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted (ℤ -> ℤ -> ℤ)%ptype
- ((projT1 args3) -> s4 -> (projT1 args))%ptype
- with
- | Some (_, (_, _))%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (ℤ -> ℤ -> ℤ)%ptype
- ((projT1 args3) -> s4 -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args3);
- v <- type.try_make_transport_cps s4 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x6 <- (if
- ((let (x6, _) := xv0 in x6) =?
- 2
- ^ (2 *
- Z.log2_up
- (let (x6, _) := xv0 in x6) / 2) -
- 1) &&
- (ZRange.normalize args0 &'
- ZRange.normalize
- (ZRange.constant
- (let (x6, _) := xv0 in x6)) <=?
- ZRange.normalize args2)%zrange
- then
- x6 <- invert_low
- (2 *
- Z.log2_up
- (let (x6, _) := xv0 in x6))
- (let (x6, _) := xv in x6);
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let (x7, _) := xv0 in x7)))%expr @
- ((##x6)%expr,
- #(Z_cast args0)%expr @
- v (Compile.reflect x5))))%expr_pat
- else None);
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end);;
- args <- invert_bind_args idc4 Raw.ident.Literal;
- args0 <- invert_bind_args idc3 Raw.ident.Z_cast;
- _ <- invert_bind_args idc2 Raw.ident.Z_land;
- args2 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted (ℤ -> ℤ -> ℤ)%ptype
- ((projT1 args3) -> s4 -> (projT1 args))%ptype
- with
- | Some (_, (_, _))%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (ℤ -> ℤ -> ℤ)%ptype
- ((projT1 args3) -> s4 -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args3);
- v <- type.try_make_transport_cps s4 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x6 <- (if
- ((let (x6, _) := xv0 in x6) =?
- 2
- ^ (2 *
- Z.log2_up
- (let (x6, _) := xv0 in x6) / 2) -
- 1) &&
- (ZRange.normalize args0 &'
- ZRange.normalize
- (ZRange.constant
- (let (x6, _) := xv0 in x6)) <=?
- ZRange.normalize args2)%zrange
- then
- x6 <- invert_high
- (2 *
- Z.log2_up
- (let (x6, _) := xv0 in x6))
- (let (x6, _) := xv in x6);
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mulhl
- (2 *
- Z.log2_up
- (let (x7, _) := xv0 in x7)))%expr @
- ((##x6)%expr,
- #(Z_cast args0)%expr @
- v (Compile.reflect x5))))%expr_pat
- else None);
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _ (_ @ _)%expr_pat
- _ | @expr.App _ _ _ s4 _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _ _ s5 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s5 _ (_ @ _)%expr_pat
+ _ | @expr.App _ _ _ s5 _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end;;
- match x4 with
- | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t3 idc3) x5 =>
- match x3 with
- | @expr.Ident _ _ _ t4 idc4 =>
- (args <- invert_bind_args idc4 Raw.ident.Literal;
- args0 <- invert_bind_args idc3 Raw.ident.Z_cast;
- _ <- invert_bind_args idc2 Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc0 Raw.ident.Literal;
+ match x5 with
+ | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t4 idc4) x6 =>
+ match x4 with
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.Ident _ _ _ t6 idc6)%expr_pat =>
+ (args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc3 Raw.ident.Z_shiftr;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc1 Raw.ident.Literal;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_mul;
match
pattern.type.unify_extracted (ℤ -> ℤ -> ℤ)%ptype
- ((projT1 args3) -> s4 -> (projT1 args))%ptype
+ ((projT1 args4) -> s5 -> (projT1 args))%ptype
with
| Some (_, (_, _))%zrange =>
if
type.type_beq base.type base.type.type_beq
(ℤ -> ℤ -> ℤ)%ptype
- ((projT1 args3) -> s4 -> (projT1 args))%ptype
+ ((projT1 args4) -> s5 -> (projT1 args))%ptype
then
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args3);
- v <- type.try_make_transport_cps s4 ℤ;
+ ##(projT2 args4);
+ v <- type.try_make_transport_cps s5 ℤ;
xv0 <- ident.unify pattern.ident.Literal
##(projT2 args);
- fv <- (x6 <- (if
- (ZRange.normalize args0 >>
+ fv <- (x8 <- (if
+ (ZRange.normalize args1 >>
ZRange.normalize
(ZRange.constant
- (let (x6, _) := xv0 in x6)) <=?
- ZRange.normalize args2)%zrange
+ (let (x8, _) := xv0 in x8)) <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args5) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args0)
then
- x6 <- invert_low
+ x8 <- invert_low
(2 *
- (let (x6, _) := xv0 in x6))
- (let (x6, _) := xv in x6);
+ (let (x8, _) := xv0 in x8))
+ (let (x8, _) := xv in x8);
Some
(#(Z_cast range)%expr @
(#(fancy_mullh
(2 *
- (let (x7, _) := xv0 in x7)))%expr @
- ((##x6)%expr,
- #(Z_cast args0)%expr @
- v (Compile.reflect x5))))%expr_pat
+ (let (x9, _) := xv0 in x9)))%expr @
+ ((##x8)%expr,
+ #(Z_cast args1)%expr @
+ v (Compile.reflect x6))))%expr_pat
else None);
- Some (Base x6));
+ Some (Base x8));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end);;
- args <- invert_bind_args idc4 Raw.ident.Literal;
- args0 <- invert_bind_args idc3 Raw.ident.Z_cast;
- _ <- invert_bind_args idc2 Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc0 Raw.ident.Literal;
+ args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc3 Raw.ident.Z_shiftr;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc1 Raw.ident.Literal;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_mul;
match
pattern.type.unify_extracted (ℤ -> ℤ -> ℤ)%ptype
- ((projT1 args3) -> s4 -> (projT1 args))%ptype
+ ((projT1 args4) -> s5 -> (projT1 args))%ptype
with
| Some (_, (_, _))%zrange =>
if
type.type_beq base.type base.type.type_beq
(ℤ -> ℤ -> ℤ)%ptype
- ((projT1 args3) -> s4 -> (projT1 args))%ptype
+ ((projT1 args4) -> s5 -> (projT1 args))%ptype
then
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args3);
- v <- type.try_make_transport_cps s4 ℤ;
+ ##(projT2 args4);
+ v <- type.try_make_transport_cps s5 ℤ;
xv0 <- ident.unify pattern.ident.Literal
##(projT2 args);
- fv <- (x6 <- (if
- (ZRange.normalize args0 >>
+ fv <- (x8 <- (if
+ (ZRange.normalize args1 >>
ZRange.normalize
(ZRange.constant
- (let (x6, _) := xv0 in x6)) <=?
- ZRange.normalize args2)%zrange
+ (let (x8, _) := xv0 in x8)) <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args5) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args0)
then
- x6 <- invert_high
- (2 * (let (x6, _) := xv0 in x6))
- (let (x6, _) := xv in x6);
+ x8 <- invert_high
+ (2 * (let (x8, _) := xv0 in x8))
+ (let (x8, _) := xv in x8);
Some
(#(Z_cast range)%expr @
(#(fancy_mulhh
- (2 * (let (x7, _) := xv0 in x7)))%expr @
- ((##x6)%expr,
- #(Z_cast args0)%expr @
- v (Compile.reflect x5))))%expr_pat
+ (2 * (let (x9, _) := xv0 in x9)))%expr @
+ ((##x8)%expr,
+ #(Z_cast args1)%expr @
+ v (Compile.reflect x6))))%expr_pat
else None);
- Some (Base x6));
+ Some (Base x8));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
+ | (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end
- | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _ (_ @ _)%expr_pat
- _ | @expr.App _ _ _ s4 _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _ _ s5 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s5 _ (_ @ _)%expr_pat
+ _ | @expr.App _ _ _ s5 _ (@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
| @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.Ident _ _ _ t1 idc1 @ #(_))%expr_pat | @expr.App _ _ _ s _
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Ident _ _ _ t2 idc2 @ #(_))%expr_pat | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.Ident _ _ _ t1 idc1 @ ($_)%expr)%expr_pat | @expr.App _ _ _ s
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Ident _ _ _ t2 idc2 @ ($_)%expr)%expr_pat | @expr.App _ _ _ s
_
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.Ident _ _ _ t1 idc1 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
@expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.Ident _ _ _ t1 idc1 @ (#(_) @ _))%expr_pat | @expr.App _ _ _ s
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Ident _ _ _ t2 idc2 @ (#(_) @ _))%expr_pat | @expr.App _ _ _ s
_
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.Ident _ _ _ t1 idc1 @ (($_)%expr @ _))%expr_pat | @expr.App _
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Ident _ _ _ t2 idc2 @ (($_)%expr @ _))%expr_pat | @expr.App _
_ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.Ident _ _ _ t1 idc1 @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Ident _ _ _ t2 idc2 @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
@expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.Ident _ _ _ t1 idc1 @ (($_)%expr @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Ident _ _ _ t2 idc2 @ (($_)%expr @ _ @ _))%expr_pat |
@expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.Ident _ _ _ t1 idc1 @ (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Ident _ _ _ t2 idc2 @ (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
@expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.Ident _ _ _ t1 idc1 @ (_ @ _ @ _ @ _))%expr_pat | @expr.App _
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Ident _ _ _ t2 idc2 @ (_ @ _ @ _ @ _))%expr_pat | @expr.App _
_ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.Ident _ _ _ t1 idc1 @ (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Ident _ _ _ t2 idc2 @ (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
@expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.Ident _ _ _ t1 idc1 @ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Ident _ _ _ t2 idc2 @ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
@expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.Ident _ _ _ t1 idc1 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
None
| @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ #(_)%expr_pat | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ ($_)%expr | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Abs _ _ _ _ _ _) | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (($_)%expr @ _)%expr_pat | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (_ @ _ @ _)%expr_pat | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.LetIn _ _ _ _ _ _ _ @ _)%expr_pat | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat)
+ (@expr.LetIn _ _ _ _ _ _ _) => None
+ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
(@expr.Ident _ _ _ t0 idc0 @ (@expr.Ident _ _ _ t1 idc1 @ x4 @ x3))%expr_pat)
x0 =>
match x4 with
- | @expr.Ident _ _ _ t2 idc2 =>
- match x3 with
- | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t3 idc3) x5 =>
- match x0 with
- | @expr.Ident _ _ _ t4 idc4 =>
- (args <- invert_bind_args idc4 Raw.ident.Literal;
- args0 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc2 Raw.ident.Literal;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s4) -> (projT1 args))%ptype
- with
- | Some (_, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
+ | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5 =>
+ match x5 with
+ | @expr.Ident _ _ _ t3 idc3 =>
+ match x3 with
+ | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t4 idc4) x6 =>
+ match x0 with
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.Ident _ _ _ t6
+ idc6)%expr_pat =>
+ args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc3 Raw.ident.Literal;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc1 Raw.ident.Z_land;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted
((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s4) -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s4 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x6 <- (if
- ((let (x6, _) := xv in x6) =?
- 2
- ^ (2 *
- Z.log2_up
- (let (x6, _) := xv in x6) /
- 2) - 1) &&
- (ZRange.normalize
- (ZRange.constant
- (let (x6, _) := xv in x6)) &'
- ZRange.normalize args0 <=?
- ZRange.normalize args3)%zrange
- then
- y <- invert_low
- (2 *
- Z.log2_up
- (let (x6, _) := xv in x6))
- (let (x6, _) := xv0 in x6);
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let (x6, _) := xv in
- x6)))%expr @
- (#(Z_cast args0)%expr @
- v (Compile.reflect x5),
- (##y)%expr)))%expr_pat
- else None);
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end);;
- args <- invert_bind_args idc4 Raw.ident.Literal;
- args0 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc2 Raw.ident.Literal;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s4) -> (projT1 args))%ptype
- with
- | Some (_, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s4) -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s4 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x6 <- (if
- ((let (x6, _) := xv in x6) =?
- 2
- ^ (2 *
- Z.log2_up
- (let (x6, _) := xv in x6) / 2) -
- 1) &&
- (ZRange.normalize
- (ZRange.constant
- (let (x6, _) := xv in x6)) &'
- ZRange.normalize args0 <=?
- ZRange.normalize args3)%zrange
- then
- y <- invert_high
- (2 *
+ (((projT1 args2) -> s5) -> (projT1 args))%ptype
+ with
+ | Some (_, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args2) -> s5) -> (projT1 args))%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args2);
+ v <- type.try_make_transport_cps s5 ℤ;
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (x8 <- (if
+ ((let (x8, _) := xv in x8) =?
+ 2
+ ^ (2 *
Z.log2_up
- (let (x6, _) := xv in x6))
- (let (x6, _) := xv0 in x6);
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mullh
- (2 *
- Z.log2_up
- (let (x6, _) := xv in x6)))%expr @
- (#(Z_cast args0)%expr @
- v (Compile.reflect x5),
- (##y)%expr)))%expr_pat
- else None);
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ x8 @ x7))%expr_pat =>
- match x8 with
- | @expr.Ident _ _ _ t6 idc6 =>
- match x7 with
- | @expr.App _ _ _ s8 _ (@expr.Ident _ _ _ t7 idc7)
- x9 =>
- args <- invert_bind_args idc7 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc6
- Raw.ident.Literal;
- _ <- invert_bind_args idc5 Raw.ident.Z_land;
- args2 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args4 <- invert_bind_args idc2
- Raw.ident.Literal;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args6 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> s4) ->
- (projT1 args0) -> s8)%ptype
- with
- | Some (_, _, (_, _))%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> s4) ->
- (projT1 args0) -> s8)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args4);
- v <- type.try_make_transport_cps s4 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v0 <- type.try_make_transport_cps s8 ℤ;
- fv <- (x10 <- (if
- ((let (x10, _) := xv in
- x10) =?
- 2
- ^ (2 *
+ (let (x8, _) := xv in x8) /
+ 2) - 1) &&
+ (ZRange.normalize
+ (ZRange.constant
+ (let (x8, _) := xv in x8)) &'
+ ZRange.normalize args1 <=?
+ ZRange.normalize args5)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args0) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args3)
+ then
+ y <- invert_low
+ (2 *
+ Z.log2_up
+ (let (x8, _) := xv in
+ x8))
+ (let (x8, _) := xv0 in
+ x8);
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mulll
+ (2 *
Z.log2_up
- (let (x10, _) :=
+ (let (x8, _) :=
xv in
- x10) / 2) - 1) &&
- ((let (x10, _) := xv0 in
- x10) =?
- 2
- ^ (2 *
+ x8)))%expr @
+ (#(Z_cast args1)%expr @
+ v (Compile.reflect x6),
+ (##y)%expr)))%expr_pat
+ else None);
+ Some (Base x8));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _
+ _)%expr_pat => None
+ | _ => None
+ end
+ | @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _ _ s5 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s5 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s5 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
+ end
+ | _ => None
+ end;;
+ match x3 with
+ | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6 =>
+ match x6 with
+ | @expr.Ident _ _ _ t4 idc4 =>
+ match x0 with
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.Ident _ _ _ t6
+ idc6)%expr_pat =>
+ args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Literal;
+ args2 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc1 Raw.ident.Z_land;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ ((s4 -> (projT1 args1)) -> (projT1 args))%ptype
+ with
+ | Some (_, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ ((s4 -> (projT1 args1)) -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s4 ℤ;
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (x8 <- (if
+ ((let (x8, _) := xv in x8) =?
+ 2
+ ^ (2 *
+ Z.log2_up
+ (let (x8, _) := xv in x8) /
+ 2) - 1) &&
+ (ZRange.normalize args3 &'
+ ZRange.normalize
+ (ZRange.constant
+ (let (x8, _) := xv in x8)) <=?
+ ZRange.normalize args5)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args0) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args2)
+ then
+ y <- invert_low
+ (2 *
+ Z.log2_up
+ (let (x8, _) := xv in
+ x8))
+ (let (x8, _) := xv0 in
+ x8);
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mulll
+ (2 *
Z.log2_up
- (let (x10, _) :=
+ (let (x8, _) :=
xv in
- x10) / 2) - 1) &&
- (ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
- xv in
- x10)) &'
- ZRange.normalize args3 <=?
- ZRange.normalize args6)%zrange &&
- (ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
- xv0 in
- x10)) &'
- ZRange.normalize args <=?
- ZRange.normalize args2)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let
- (x10, _) :=
- xv in
- x10)))%expr @
- (#(Z_cast args3)%expr @
- v
- (Compile.reflect x5),
- #(Z_cast args)%expr @
- v0
- (Compile.reflect x9))))%expr_pat
- else None);
- Some (Base x10));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s8 _ ($_)%expr _ | @expr.App _ _
- _ s8 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
- _ s8 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s8 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
+ x8)))%expr @
+ (#(Z_cast args3)%expr @
+ v (Compile.reflect x5),
+ (##y)%expr)))%expr_pat
+ else None);
+ Some (Base x8));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
end
- | @expr.App _ _ _ s8 _ (@expr.Ident _ _ _ t6 idc6)
- x9 =>
- match x7 with
- | @expr.Ident _ _ _ t7 idc7 =>
- args <- invert_bind_args idc7 Raw.ident.Literal;
- args0 <- invert_bind_args idc6 Raw.ident.Z_cast;
- _ <- invert_bind_args idc5 Raw.ident.Z_land;
- args2 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args4 <- invert_bind_args idc2
- Raw.ident.Literal;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args6 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> s4) ->
- s8 -> (projT1 args))%ptype
- with
- | Some (_, _, (_, _))%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> s4) ->
- s8 -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args4);
- v <- type.try_make_transport_cps s4 ℤ;
- v0 <- type.try_make_transport_cps s8 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x10 <- (if
- ((let (x10, _) := xv in
- x10) =?
- 2
- ^ (2 *
+ | (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _
+ _)%expr_pat => None
+ | _ => None
+ end
+ | _ => None
+ end;;
+ match x5 with
+ | @expr.Ident _ _ _ t4 idc4 =>
+ match x0 with
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.Ident _ _ _ t6
+ idc6)%expr_pat =>
+ args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Literal;
+ args2 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc1 Raw.ident.Z_land;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> s5) -> (projT1 args))%ptype
+ with
+ | Some (_, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ (((projT1 args1) -> s5) -> (projT1 args))%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ v <- type.try_make_transport_cps s5 ℤ;
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (x8 <- (if
+ ((let (x8, _) := xv in x8) =?
+ 2
+ ^ (2 *
+ Z.log2_up
+ (let (x8, _) := xv in x8) /
+ 2) - 1) &&
+ (ZRange.normalize
+ (ZRange.constant
+ (let (x8, _) := xv in x8)) &'
+ ZRange.normalize args2 <=?
+ ZRange.normalize args5)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args0) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args3)
+ then
+ y <- invert_high
+ (2 *
+ Z.log2_up
+ (let (x8, _) := xv in
+ x8))
+ (let (x8, _) := xv0 in
+ x8);
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mullh
+ (2 *
Z.log2_up
- (let (x10, _) :=
+ (let (x8, _) :=
xv in
- x10) / 2) - 1) &&
- ((let (x10, _) := xv0 in
- x10) =?
- 2
- ^ (2 *
+ x8)))%expr @
+ (#(Z_cast args2)%expr @
+ v (Compile.reflect x6),
+ (##y)%expr)))%expr_pat
+ else None);
+ Some (Base x8));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _
+ _)%expr_pat => None
+ | _ => None
+ end
+ | _ => None
+ end;;
+ match x6 with
+ | @expr.Ident _ _ _ t4 idc4 =>
+ match x0 with
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.Ident _ _ _ t6
+ idc6)%expr_pat =>
+ args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Literal;
+ args2 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc1 Raw.ident.Z_land;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ ((s4 -> (projT1 args1)) -> (projT1 args))%ptype
+ with
+ | Some (_, _, _)%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ)%ptype
+ ((s4 -> (projT1 args1)) -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s4 ℤ;
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args1);
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (x8 <- (if
+ ((let (x8, _) := xv in x8) =?
+ 2
+ ^ (2 *
+ Z.log2_up
+ (let (x8, _) := xv in x8) /
+ 2) - 1) &&
+ (ZRange.normalize args3 &'
+ ZRange.normalize
+ (ZRange.constant
+ (let (x8, _) := xv in x8)) <=?
+ ZRange.normalize args5)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args0) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args2)
+ then
+ y <- invert_high
+ (2 *
+ Z.log2_up
+ (let (x8, _) := xv in
+ x8))
+ (let (x8, _) := xv0 in
+ x8);
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mullh
+ (2 *
Z.log2_up
- (let (x10, _) :=
- xv in
- x10) / 2) - 1) &&
- (ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
+ (let (x8, _) :=
xv in
- x10)) &'
- ZRange.normalize args3 <=?
- ZRange.normalize args6)%zrange &&
- (ZRange.normalize args0 &'
- ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
- xv0 in
- x10)) <=?
- ZRange.normalize args2)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let
- (x10, _) :=
- xv in
- x10)))%expr @
- (#(Z_cast args3)%expr @
- v
- (Compile.reflect x5),
- #(Z_cast args0)%expr @
- v0
- (Compile.reflect x9))))%expr_pat
- else None);
- Some (Base x10));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
+ x8)))%expr @
+ (#(Z_cast args3)%expr @
+ v (Compile.reflect x5),
+ (##y)%expr)))%expr_pat
+ else None);
+ Some (Base x8));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
end
- | @expr.App _ _ _ s8 _ ($_)%expr _ | @expr.App _ _ _ s8
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s8 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s8 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _
+ _)%expr_pat => None
| _ => None
- end;;
- match x8 with
- | @expr.App _ _ _ s8 _ (@expr.Ident _ _ _ t6 idc6)
- x9 =>
- match x7 with
- | @expr.Ident _ _ _ t7 idc7 =>
- args <- invert_bind_args idc7 Raw.ident.Literal;
- args0 <- invert_bind_args idc6 Raw.ident.Z_cast;
- _ <- invert_bind_args idc5 Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args4 <- invert_bind_args idc2
- Raw.ident.Literal;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args6 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> s4) ->
- s8 -> (projT1 args))%ptype
- with
- | Some (_, _, (_, _))%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> s4) ->
- s8 -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args4);
- v <- type.try_make_transport_cps s4 ℤ;
- v0 <- type.try_make_transport_cps s8 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x10 <- (if
- ((let (x10, _) := xv in
- x10) =?
- 2
- ^ (2 *
- (let (x10, _) :=
- xv0 in
- x10) / 2) - 1) &&
- (ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
- xv in
- x10)) &'
- ZRange.normalize args3 <=?
- ZRange.normalize args6)%zrange &&
- (ZRange.normalize args0 >>
- ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
- xv0 in
- x10)) <=?
- ZRange.normalize args2)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mullh
- (2 *
- (let (x10, _) :=
- xv0 in
- x10)))%expr @
- (#(Z_cast args3)%expr @
- v
- (Compile.reflect x5),
- #(Z_cast args0)%expr @
- v0
- (Compile.reflect x9))))%expr_pat
- else None);
- Some (Base x10));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
+ end
+ | _ => None
+ end;;
+ match x5 with
+ | @expr.Ident _ _ _ t4 idc4 =>
+ match x0 with
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t9 idc9) x11))%expr_pat =>
+ args <- invert_bind_args idc9 Raw.ident.Z_cast;
+ args0 <- invert_bind_args idc8 Raw.ident.Literal;
+ args1 <- invert_bind_args idc7 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc6 Raw.ident.Z_land;
+ args3 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc4 Raw.ident.Literal;
+ args5 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args6 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc1 Raw.ident.Z_land;
+ args8 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ (((projT1 args4) -> s5) ->
+ (projT1 args0) -> s10)%ptype
+ with
+ | Some (_, _, (_, _))%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ (((projT1 args4) -> s5) ->
+ (projT1 args0) -> s10)%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args4);
+ v <- type.try_make_transport_cps s5 ℤ;
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ v0 <- type.try_make_transport_cps s10 ℤ;
+ fv <- (x12 <- (if
+ ((let (x12, _) := xv in x12) =?
+ 2
+ ^ (2 *
+ Z.log2_up
+ (let (x12, _) := xv in
+ x12) / 2) - 1) &&
+ ((let (x12, _) := xv0 in x12) =?
+ 2
+ ^ (2 *
+ Z.log2_up
+ (let (x12, _) := xv in
+ x12) / 2) - 1) &&
+ (ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) := xv in
+ x12)) &'
+ ZRange.normalize args5 <=?
+ ZRange.normalize args8)%zrange &&
+ (ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) := xv0 in
+ x12)) &'
+ ZRange.normalize args <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x12, _) := xv in x12)
+ (ZRange.normalize args6) &&
+ is_bounded_by_bool
+ (let (x12, _) := xv0 in x12)
+ (ZRange.normalize args1)
+ then
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mulll
+ (2 *
+ Z.log2_up
+ (let (x12, _) :=
+ xv in
+ x12)))%expr @
+ (#(Z_cast args5)%expr @
+ v (Compile.reflect x6),
+ #(Z_cast args)%expr @
+ v0 (Compile.reflect x11))))%expr_pat
+ else None);
+ Some (Base x12));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
end
- | @expr.App _ _ _ s8 _ ($_)%expr _ | @expr.App _ _ _ s8
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s8 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s8 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.App _ _ _ s10 _ ($_)%expr _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.App _ _ _ s10 _
+ (@expr.Abs _ _ _ _ _ _) _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.App _ _ _ s10 _ (_ @ _) _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.App _ _ _ s10 _
+ (@expr.LetIn _ _ _ _ _ _ _) _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ #(_)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.Abs _ _ _ _ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
+ None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ ($_)%expr) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Abs _ _ _ _ _ _) @
+ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ (_ @ _)) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.LetIn _ _ _ _ _
+ _ _) @ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _ _ _ @
+ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (($_)%expr @ _) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Abs _ _ _ _ _ _ @ _) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (_ @ _ @ _) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _ _ _ _
+ _ @ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _
+ _)%expr_pat => None
| _ => None
end
- | (@expr.Ident _ _ _ t4 idc4 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
- None
| _ => None
- end
- | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s4 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5 =>
- match x3 with
- | @expr.Ident _ _ _ t3 idc3 =>
- match x0 with
+ end;;
+ match x6 with
| @expr.Ident _ _ _ t4 idc4 =>
- (args <- invert_bind_args idc4 Raw.ident.Literal;
- args0 <- invert_bind_args idc3 Raw.ident.Literal;
- args1 <- invert_bind_args idc2 Raw.ident.Z_cast;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- ((s4 -> (projT1 args0)) -> (projT1 args))%ptype
- with
- | Some (_, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ)%ptype
- ((s4 -> (projT1 args0)) -> (projT1 args))%ptype
- then
- v <- type.try_make_transport_cps s4 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x6 <- (if
- ((let (x6, _) := xv in x6) =?
- 2
- ^ (2 *
- Z.log2_up
- (let (x6, _) := xv in x6) /
- 2) - 1) &&
- (ZRange.normalize args1 &'
- ZRange.normalize
- (ZRange.constant
- (let (x6, _) := xv in x6)) <=?
- ZRange.normalize args3)%zrange
- then
- y <- invert_low
- (2 *
+ match x0 with
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t9 idc9) x11))%expr_pat =>
+ args <- invert_bind_args idc9 Raw.ident.Z_cast;
+ args0 <- invert_bind_args idc8 Raw.ident.Literal;
+ args1 <- invert_bind_args idc7 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc6 Raw.ident.Z_land;
+ args3 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc4 Raw.ident.Literal;
+ args5 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args6 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc1 Raw.ident.Z_land;
+ args8 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s4 -> (projT1 args4)) ->
+ (projT1 args0) -> s10)%ptype
+ with
+ | Some (_, _, (_, _))%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s4 -> (projT1 args4)) ->
+ (projT1 args0) -> s10)%ptype
+ then
+ v <- type.try_make_transport_cps s4 ℤ;
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args4);
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args0);
+ v0 <- type.try_make_transport_cps s10 ℤ;
+ fv <- (x12 <- (if
+ ((let (x12, _) := xv in x12) =?
+ 2
+ ^ (2 *
Z.log2_up
- (let (x6, _) := xv in x6))
- (let (x6, _) := xv0 in x6);
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let (x6, _) := xv in
- x6)))%expr @
- (#(Z_cast args1)%expr @
- v (Compile.reflect x5),
- (##y)%expr)))%expr_pat
- else None);
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end);;
- args <- invert_bind_args idc4 Raw.ident.Literal;
- args0 <- invert_bind_args idc3 Raw.ident.Literal;
- args1 <- invert_bind_args idc2 Raw.ident.Z_cast;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- ((s4 -> (projT1 args0)) -> (projT1 args))%ptype
- with
- | Some (_, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ)%ptype
- ((s4 -> (projT1 args0)) -> (projT1 args))%ptype
- then
- v <- type.try_make_transport_cps s4 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x6 <- (if
- ((let (x6, _) := xv in x6) =?
- 2
- ^ (2 *
- Z.log2_up
- (let (x6, _) := xv in x6) / 2) -
- 1) &&
- (ZRange.normalize args1 &'
- ZRange.normalize
- (ZRange.constant
- (let (x6, _) := xv in x6)) <=?
- ZRange.normalize args3)%zrange
- then
- y <- invert_high
- (2 *
- Z.log2_up
- (let (x6, _) := xv in x6))
- (let (x6, _) := xv0 in x6);
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mullh
- (2 *
+ (let (x12, _) := xv in
+ x12) / 2) - 1) &&
+ ((let (x12, _) := xv0 in x12) =?
+ 2
+ ^ (2 *
Z.log2_up
- (let (x6, _) := xv in x6)))%expr @
- (#(Z_cast args1)%expr @
- v (Compile.reflect x5),
- (##y)%expr)))%expr_pat
- else None);
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
+ (let (x12, _) := xv in
+ x12) / 2) - 1) &&
+ (ZRange.normalize args6 &'
+ ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) := xv in
+ x12)) <=?
+ ZRange.normalize args8)%zrange &&
+ (ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) := xv0 in
+ x12)) &'
+ ZRange.normalize args <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x12, _) := xv in x12)
+ (ZRange.normalize args5) &&
+ is_bounded_by_bool
+ (let (x12, _) := xv0 in x12)
+ (ZRange.normalize args1)
+ then
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mulll
+ (2 *
+ Z.log2_up
+ (let (x12, _) :=
+ xv in
+ x12)))%expr @
+ (#(Z_cast args6)%expr @
+ v (Compile.reflect x5),
+ #(Z_cast args)%expr @
+ v0 (Compile.reflect x11))))%expr_pat
+ else None);
+ Some (Base x12));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.App _ _ _ s10 _ ($_)%expr _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.App _ _ _ s10 _
+ (@expr.Abs _ _ _ _ _ _) _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.App _ _ _ s10 _ (_ @ _) _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.App _ _ _ s10 _
+ (@expr.LetIn _ _ _ _ _ _ _) _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ #(_)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.Abs _ _ _ _ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _ t8
+ idc8) @ @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
+ None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ ($_)%expr) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Abs _ _ _ _ _ _) @
+ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ (_ @ _)) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.LetIn _ _ _ _ _
+ _ _) @ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _ _ _ @
+ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (($_)%expr @ _) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Abs _ _ _ _ _ _ @ _) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (_ @ _ @ _) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _ _ _ _
+ _ @ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _
+ _)%expr_pat => None
+ | _ => None
end
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ x8 @ x7))%expr_pat =>
- match x8 with
- | @expr.Ident _ _ _ t6 idc6 =>
- match x7 with
- | @expr.App _ _ _ s8 _ (@expr.Ident _ _ _ t7 idc7)
- x9 =>
- args <- invert_bind_args idc7 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc6
- Raw.ident.Literal;
- _ <- invert_bind_args idc5 Raw.ident.Z_land;
- args2 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc3
- Raw.ident.Literal;
- args4 <- invert_bind_args idc2 Raw.ident.Z_cast;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args6 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- ((s4 -> (projT1 args3)) ->
- (projT1 args0) -> s8)%ptype
- with
- | Some (_, _, (_, _))%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- ((s4 -> (projT1 args3)) ->
- (projT1 args0) -> s8)%ptype
- then
- v <- type.try_make_transport_cps s4 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args3);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v0 <- type.try_make_transport_cps s8 ℤ;
- fv <- (x10 <- (if
- ((let (x10, _) := xv in
- x10) =?
- 2
- ^ (2 *
- Z.log2_up
- (let (x10, _) :=
- xv in
- x10) / 2) - 1) &&
- ((let (x10, _) := xv0 in
- x10) =?
- 2
- ^ (2 *
- Z.log2_up
- (let (x10, _) :=
- xv in
- x10) / 2) - 1) &&
- (ZRange.normalize args4 &'
- ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
- xv in
- x10)) <=?
- ZRange.normalize args6)%zrange &&
- (ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
- xv0 in
- x10)) &'
- ZRange.normalize args <=?
- ZRange.normalize args2)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let
- (x10, _) :=
- xv in
- x10)))%expr @
- (#(Z_cast args4)%expr @
- v
- (Compile.reflect x5),
- #(Z_cast args)%expr @
- v0
- (Compile.reflect x9))))%expr_pat
- else None);
- Some (Base x10));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s8 _ ($_)%expr _ | @expr.App _ _
- _ s8 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
- _ s8 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s8 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
+ | _ => None
+ end;;
+ match x5 with
+ | @expr.Ident _ _ _ t4 idc4 =>
+ match x0 with
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9)))%expr_pat =>
+ args <- invert_bind_args idc9 Raw.ident.Literal;
+ args0 <- invert_bind_args idc8 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc7 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc6 Raw.ident.Z_land;
+ args3 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc4 Raw.ident.Literal;
+ args5 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args6 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc1 Raw.ident.Z_land;
+ args8 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ (((projT1 args4) -> s5) -> s9 -> (projT1 args))%ptype
+ with
+ | Some (_, _, (_, _))%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ (((projT1 args4) -> s5) ->
+ s9 -> (projT1 args))%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args4);
+ v <- type.try_make_transport_cps s5 ℤ;
+ v0 <- type.try_make_transport_cps s9 ℤ;
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (x12 <- (if
+ ((let (x12, _) := xv in x12) =?
+ 2
+ ^ (2 *
+ Z.log2_up
+ (let (x12, _) := xv in
+ x12) / 2) - 1) &&
+ ((let (x12, _) := xv0 in x12) =?
+ 2
+ ^ (2 *
+ Z.log2_up
+ (let (x12, _) := xv in
+ x12) / 2) - 1) &&
+ (ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) := xv in
+ x12)) &'
+ ZRange.normalize args5 <=?
+ ZRange.normalize args8)%zrange &&
+ (ZRange.normalize args1 &'
+ ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) := xv0 in
+ x12)) <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x12, _) := xv in x12)
+ (ZRange.normalize args6) &&
+ is_bounded_by_bool
+ (let (x12, _) := xv0 in x12)
+ (ZRange.normalize args0)
+ then
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mulll
+ (2 *
+ Z.log2_up
+ (let (x12, _) :=
+ xv in
+ x12)))%expr @
+ (#(Z_cast args5)%expr @
+ v (Compile.reflect x6),
+ #(Z_cast args1)%expr @
+ v0 (Compile.reflect x10))))%expr_pat
+ else None);
+ Some (Base x12));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
end
- | @expr.App _ _ _ s8 _ (@expr.Ident _ _ _ t6 idc6)
- x9 =>
- match x7 with
- | @expr.Ident _ _ _ t7 idc7 =>
- args <- invert_bind_args idc7 Raw.ident.Literal;
- args0 <- invert_bind_args idc6 Raw.ident.Z_cast;
- _ <- invert_bind_args idc5 Raw.ident.Z_land;
- args2 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc3
- Raw.ident.Literal;
- args4 <- invert_bind_args idc2 Raw.ident.Z_cast;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args6 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- ((s4 -> (projT1 args3)) ->
- s8 -> (projT1 args))%ptype
- with
- | Some (_, _, (_, _))%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- ((s4 -> (projT1 args3)) ->
- s8 -> (projT1 args))%ptype
- then
- v <- type.try_make_transport_cps s4 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args3);
- v0 <- type.try_make_transport_cps s8 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x10 <- (if
- ((let (x10, _) := xv in
- x10) =?
- 2
- ^ (2 *
- Z.log2_up
- (let (x10, _) :=
- xv in
- x10) / 2) - 1) &&
- ((let (x10, _) := xv0 in
- x10) =?
- 2
- ^ (2 *
- Z.log2_up
- (let (x10, _) :=
- xv in
- x10) / 2) - 1) &&
- (ZRange.normalize args4 &'
- ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
- xv in
- x10)) <=?
- ZRange.normalize args6)%zrange &&
- (ZRange.normalize args0 &'
- ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
- xv0 in
- x10)) <=?
- ZRange.normalize args2)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mulll
- (2 *
- Z.log2_up
- (let
- (x10, _) :=
- xv in
- x10)))%expr @
- (#(Z_cast args4)%expr @
- v
- (Compile.reflect x5),
- #(Z_cast args0)%expr @
- v0
- (Compile.reflect x9))))%expr_pat
- else None);
- Some (Base x10));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ ($_)%expr)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Abs _ _ _ _ _ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ (_ @ _))))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.LetIn _ _ _ _ _
+ _ _)))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ #(_)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ @expr.Abs _ _ _ _
+ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ (($_)%expr @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Abs _ _ _ _ _ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ (_ @ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ @expr.LetIn _ _ _
+ _ _ _ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ ($_)%expr _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (_ @ _) _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
+ None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _ _ _ @
+ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _ _ _ _
+ _ @ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _
+ _)%expr_pat => None
+ | _ => None
+ end
+ | _ => None
+ end;;
+ match x6 with
+ | @expr.Ident _ _ _ t4 idc4 =>
+ match x0 with
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9)))%expr_pat =>
+ args <- invert_bind_args idc9 Raw.ident.Literal;
+ args0 <- invert_bind_args idc8 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc7 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc6 Raw.ident.Z_land;
+ args3 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc4 Raw.ident.Literal;
+ args5 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args6 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc1 Raw.ident.Z_land;
+ args8 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s4 -> (projT1 args4)) -> s9 -> (projT1 args))%ptype
+ with
+ | Some (_, _, (_, _))%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s4 -> (projT1 args4)) ->
+ s9 -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s4 ℤ;
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args4);
+ v0 <- type.try_make_transport_cps s9 ℤ;
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (x12 <- (if
+ ((let (x12, _) := xv in x12) =?
+ 2
+ ^ (2 *
+ Z.log2_up
+ (let (x12, _) := xv in
+ x12) / 2) - 1) &&
+ ((let (x12, _) := xv0 in x12) =?
+ 2
+ ^ (2 *
+ Z.log2_up
+ (let (x12, _) := xv in
+ x12) / 2) - 1) &&
+ (ZRange.normalize args6 &'
+ ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) := xv in
+ x12)) <=?
+ ZRange.normalize args8)%zrange &&
+ (ZRange.normalize args1 &'
+ ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) := xv0 in
+ x12)) <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x12, _) := xv in x12)
+ (ZRange.normalize args5) &&
+ is_bounded_by_bool
+ (let (x12, _) := xv0 in x12)
+ (ZRange.normalize args0)
+ then
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mulll
+ (2 *
+ Z.log2_up
+ (let (x12, _) :=
+ xv in
+ x12)))%expr @
+ (#(Z_cast args6)%expr @
+ v (Compile.reflect x5),
+ #(Z_cast args1)%expr @
+ v0 (Compile.reflect x10))))%expr_pat
+ else None);
+ Some (Base x12));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
end
- | @expr.App _ _ _ s8 _ ($_)%expr _ | @expr.App _ _ _ s8
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s8 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s8 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ ($_)%expr)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Abs _ _ _ _ _ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ (_ @ _))))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.LetIn _ _ _ _ _
+ _ _)))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ #(_)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ @expr.Abs _ _ _ _
+ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ (($_)%expr @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Abs _ _ _ _ _ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ (_ @ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ @expr.LetIn _ _ _
+ _ _ _ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ ($_)%expr _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (_ @ _) _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
+ None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _ _ _ @
+ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _ _ _ _
+ _ @ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _
+ _)%expr_pat => None
| _ => None
- end;;
- match x8 with
- | @expr.App _ _ _ s8 _ (@expr.Ident _ _ _ t6 idc6)
- x9 =>
- match x7 with
- | @expr.Ident _ _ _ t7 idc7 =>
- args <- invert_bind_args idc7 Raw.ident.Literal;
- args0 <- invert_bind_args idc6 Raw.ident.Z_cast;
- _ <- invert_bind_args idc5 Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc3
- Raw.ident.Literal;
- args4 <- invert_bind_args idc2 Raw.ident.Z_cast;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args6 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- ((s4 -> (projT1 args3)) ->
- s8 -> (projT1 args))%ptype
- with
- | Some (_, _, (_, _))%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- ((s4 -> (projT1 args3)) ->
- s8 -> (projT1 args))%ptype
- then
- v <- type.try_make_transport_cps s4 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args3);
- v0 <- type.try_make_transport_cps s8 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x10 <- (if
- ((let (x10, _) := xv in
- x10) =?
- 2
- ^ (2 *
- (let (x10, _) :=
- xv0 in
- x10) / 2) - 1) &&
- (ZRange.normalize args4 &'
- ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
- xv in
- x10)) <=?
- ZRange.normalize args6)%zrange &&
- (ZRange.normalize args0 >>
- ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
- xv0 in
- x10)) <=?
- ZRange.normalize args2)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mullh
- (2 *
- (let (x10, _) :=
- xv0 in
- x10)))%expr @
- (#(Z_cast args4)%expr @
- v
- (Compile.reflect x5),
- #(Z_cast args0)%expr @
- v0
- (Compile.reflect x9))))%expr_pat
- else None);
- Some (Base x10));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
+ end
+ | _ => None
+ end;;
+ match x5 with
+ | @expr.Ident _ _ _ t4 idc4 =>
+ match x0 with
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9)))%expr_pat =>
+ args <- invert_bind_args idc9 Raw.ident.Literal;
+ args0 <- invert_bind_args idc8 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc7 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc6 Raw.ident.Z_shiftr;
+ args3 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc4 Raw.ident.Literal;
+ args5 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args6 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc1 Raw.ident.Z_land;
+ args8 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ (((projT1 args4) -> s5) -> s9 -> (projT1 args))%ptype
+ with
+ | Some (_, _, (_, _))%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ (((projT1 args4) -> s5) ->
+ s9 -> (projT1 args))%ptype
+ then
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args4);
+ v <- type.try_make_transport_cps s5 ℤ;
+ v0 <- type.try_make_transport_cps s9 ℤ;
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (x12 <- (if
+ ((let (x12, _) := xv in x12) =?
+ 2
+ ^ (2 *
+ (let (x12, _) := xv0 in
+ x12) / 2) - 1) &&
+ (ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) := xv in
+ x12)) &'
+ ZRange.normalize args5 <=?
+ ZRange.normalize args8)%zrange &&
+ (ZRange.normalize args1 >>
+ ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) := xv0 in
+ x12)) <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x12, _) := xv in x12)
+ (ZRange.normalize args6) &&
+ is_bounded_by_bool
+ (let (x12, _) := xv0 in x12)
+ (ZRange.normalize args0)
+ then
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mullh
+ (2 *
+ (let (x12, _) :=
+ xv0 in
+ x12)))%expr @
+ (#(Z_cast args5)%expr @
+ v (Compile.reflect x6),
+ #(Z_cast args1)%expr @
+ v0 (Compile.reflect x10))))%expr_pat
+ else None);
+ Some (Base x12));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
end
- | @expr.App _ _ _ s8 _ ($_)%expr _ | @expr.App _ _ _ s8
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s8 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s8 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ ($_)%expr)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Abs _ _ _ _ _ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ (_ @ _))))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.LetIn _ _ _ _ _
+ _ _)))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ #(_)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ @expr.Abs _ _ _ _
+ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ (($_)%expr @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Abs _ _ _ _ _ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ (_ @ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ @expr.LetIn _ _ _
+ _ _ _ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ ($_)%expr _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (_ @ _) _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
+ None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _ _ _ @
+ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _ _ _ _
+ _ @ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _
+ _)%expr_pat => None
+ | _ => None
+ end
+ | _ => None
+ end;;
+ match x6 with
+ | @expr.Ident _ _ _ t4 idc4 =>
+ match x0 with
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9)))%expr_pat =>
+ args <- invert_bind_args idc9 Raw.ident.Literal;
+ args0 <- invert_bind_args idc8 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc7 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc6 Raw.ident.Z_shiftr;
+ args3 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc4 Raw.ident.Literal;
+ args5 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args6 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc1 Raw.ident.Z_land;
+ args8 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s4 -> (projT1 args4)) -> s9 -> (projT1 args))%ptype
+ with
+ | Some (_, _, (_, _))%zrange =>
+ if
+ type.type_beq base.type base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s4 -> (projT1 args4)) ->
+ s9 -> (projT1 args))%ptype
+ then
+ v <- type.try_make_transport_cps s4 ℤ;
+ xv <- ident.unify pattern.ident.Literal
+ ##(projT2 args4);
+ v0 <- type.try_make_transport_cps s9 ℤ;
+ xv0 <- ident.unify pattern.ident.Literal
+ ##(projT2 args);
+ fv <- (x12 <- (if
+ ((let (x12, _) := xv in x12) =?
+ 2
+ ^ (2 *
+ (let (x12, _) := xv0 in
+ x12) / 2) - 1) &&
+ (ZRange.normalize args6 &'
+ ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) := xv in
+ x12)) <=?
+ ZRange.normalize args8)%zrange &&
+ (ZRange.normalize args1 >>
+ ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) := xv0 in
+ x12)) <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x12, _) := xv in x12)
+ (ZRange.normalize args5) &&
+ is_bounded_by_bool
+ (let (x12, _) := xv0 in x12)
+ (ZRange.normalize args0)
+ then
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mullh
+ (2 *
+ (let (x12, _) :=
+ xv0 in
+ x12)))%expr @
+ (#(Z_cast args6)%expr @
+ v (Compile.reflect x5),
+ #(Z_cast args1)%expr @
+ v0 (Compile.reflect x10))))%expr_pat
+ else None);
+ Some (Base x12));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ ($_)%expr)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Abs _ _ _ _ _ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ (_ @ _))))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.LetIn _ _ _ _ _
+ _ _)))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ #(_)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ @expr.Abs _ _ _ _
+ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ (($_)%expr @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.Abs _ _ _ _ _ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ (_ @ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t7 idc7) x10 @ @expr.LetIn _ _ _
+ _ _ _ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ ($_)%expr _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (_ @ _) _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s9 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
+ None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _ _ _ @
+ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _ _ _ _
+ _ @ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _
+ _)%expr_pat => None
| _ => None
end
- | (@expr.Ident _ _ _ t4 idc4 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
- None
| _ => None
end
+ | @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _ _ s5 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s5 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s5 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
| @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4 _
@@ -1414,371 +2072,462 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
match x4 with
| @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5 =>
match x3 with
- | @expr.Ident _ _ _ t3 idc3 =>
+ | (@expr.Ident _ _ _ t3 idc3 @ @expr.Ident _ _ _ t4 idc4)%expr_pat =>
match x0 with
- | @expr.Ident _ _ _ t4 idc4 =>
- (args <- invert_bind_args idc4 Raw.ident.Literal;
- args0 <- invert_bind_args idc3 Raw.ident.Literal;
- args1 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.Ident _ _ _ t6 idc6)%expr_pat =>
+ (args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Literal;
+ args2 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
_ <- invert_bind_args idc1 Raw.ident.Z_shiftr;
- args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_mul;
match
pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- ((s4 -> (projT1 args0)) -> (projT1 args))%ptype
+ ((s4 -> (projT1 args1)) -> (projT1 args))%ptype
with
| Some (_, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
((ℤ -> ℤ) -> ℤ)%ptype
- ((s4 -> (projT1 args0)) -> (projT1 args))%ptype
+ ((s4 -> (projT1 args1)) -> (projT1 args))%ptype
then
v <- type.try_make_transport_cps s4 ℤ;
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
+ ##(projT2 args1);
xv0 <- ident.unify pattern.ident.Literal
##(projT2 args);
- fv <- (x6 <- (if
- (ZRange.normalize args1 >>
+ fv <- (x8 <- (if
+ (ZRange.normalize args3 >>
ZRange.normalize
(ZRange.constant
- (let (x6, _) := xv in x6)) <=?
- ZRange.normalize args3)%zrange
+ (let (x8, _) := xv in x8)) <=?
+ ZRange.normalize args5)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args0) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args2)
then
y <- invert_low
(2 *
- (let (x6, _) := xv in x6))
- (let (x6, _) := xv0 in x6);
+ (let (x8, _) := xv in x8))
+ (let (x8, _) := xv0 in x8);
Some
(#(Z_cast range)%expr @
(#(fancy_mulhl
(2 *
- (let (x6, _) := xv in x6)))%expr @
- (#(Z_cast args1)%expr @
+ (let (x8, _) := xv in x8)))%expr @
+ (#(Z_cast args3)%expr @
v (Compile.reflect x5),
(##y)%expr)))%expr_pat
else None);
- Some (Base x6));
+ Some (Base x8));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end);;
- args <- invert_bind_args idc4 Raw.ident.Literal;
- args0 <- invert_bind_args idc3 Raw.ident.Literal;
- args1 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args <- invert_bind_args idc6 Raw.ident.Literal;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Literal;
+ args2 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
_ <- invert_bind_args idc1 Raw.ident.Z_shiftr;
- args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_mul;
match
pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- ((s4 -> (projT1 args0)) -> (projT1 args))%ptype
+ ((s4 -> (projT1 args1)) -> (projT1 args))%ptype
with
| Some (_, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
((ℤ -> ℤ) -> ℤ)%ptype
- ((s4 -> (projT1 args0)) -> (projT1 args))%ptype
+ ((s4 -> (projT1 args1)) -> (projT1 args))%ptype
then
v <- type.try_make_transport_cps s4 ℤ;
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
+ ##(projT2 args1);
xv0 <- ident.unify pattern.ident.Literal
##(projT2 args);
- fv <- (x6 <- (if
- (ZRange.normalize args1 >>
+ fv <- (x8 <- (if
+ (ZRange.normalize args3 >>
ZRange.normalize
(ZRange.constant
- (let (x6, _) := xv in x6)) <=?
- ZRange.normalize args3)%zrange
+ (let (x8, _) := xv in x8)) <=?
+ ZRange.normalize args5)%zrange &&
+ is_bounded_by_bool
+ (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args0) &&
+ is_bounded_by_bool
+ (let (x8, _) := xv in x8)
+ (ZRange.normalize args2)
then
y <- invert_high
(2 *
- (let (x6, _) := xv in x6))
- (let (x6, _) := xv0 in x6);
+ (let (x8, _) := xv in x8))
+ (let (x8, _) := xv0 in x8);
Some
(#(Z_cast range)%expr @
(#(fancy_mulhh
(2 *
- (let (x6, _) := xv in x6)))%expr @
- (#(Z_cast args1)%expr @
+ (let (x8, _) := xv in x8)))%expr @
+ (#(Z_cast args3)%expr @
v (Compile.reflect x5),
(##y)%expr)))%expr_pat
else None);
- Some (Base x6));
+ Some (Base x8));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ x8 @ x7))%expr_pat =>
- match x8 with
- | @expr.Ident _ _ _ t6 idc6 =>
- match x7 with
- | @expr.App _ _ _ s8 _ (@expr.Ident _ _ _ t7 idc7)
- x9 =>
- args <- invert_bind_args idc7 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc6
- Raw.ident.Literal;
- _ <- invert_bind_args idc5 Raw.ident.Z_land;
- args2 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc3
- Raw.ident.Literal;
- args4 <- invert_bind_args idc2 Raw.ident.Z_cast;
- _ <- invert_bind_args idc1 Raw.ident.Z_shiftr;
- args6 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_mul;
- match
- pattern.type.unify_extracted
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- ((s4 -> (projT1 args3)) ->
- (projT1 args0) -> s8)%ptype
- with
- | Some (_, _, (_, _))%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- ((s4 -> (projT1 args3)) ->
- (projT1 args0) -> s8)%ptype
- then
- v <- type.try_make_transport_cps s4 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args3);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v0 <- type.try_make_transport_cps s8 ℤ;
- fv <- (x10 <- (if
- ((let (x10, _) := xv0 in
- x10) =?
- 2
- ^ (2 *
- (let (x10, _) := xv in
- x10) / 2) - 1) &&
- (ZRange.normalize args4 >>
- ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
- xv in
- x10)) <=?
- ZRange.normalize args6)%zrange &&
- (ZRange.normalize
- (ZRange.constant
- (let (x10, _) :=
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ x9 @ x8))%expr_pat =>
+ match x9 with
+ | @expr.App _ _ _ s9 _ (@expr.Ident _ _ _ t7 idc7)
+ x10 =>
+ match x10 with
+ | @expr.Ident _ _ _ t8 idc8 =>
+ match x8 with
+ | @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t9 idc9) x11 =>
+ args <- invert_bind_args idc9
+ Raw.ident.Z_cast;
+ args0 <- invert_bind_args idc8
+ Raw.ident.Literal;
+ args1 <- invert_bind_args idc7
+ Raw.ident.Z_cast;
+ _ <- invert_bind_args idc6 Raw.ident.Z_land;
+ args3 <- invert_bind_args idc5
+ Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc4
+ Raw.ident.Literal;
+ args5 <- invert_bind_args idc3
+ Raw.ident.Z_cast;
+ args6 <- invert_bind_args idc2
+ Raw.ident.Z_cast;
+ _ <- invert_bind_args idc1
+ Raw.ident.Z_shiftr;
+ args8 <- invert_bind_args idc0
+ Raw.ident.Z_cast;
+ _ <- invert_bind_args idc Raw.ident.Z_mul;
+ match
+ pattern.type.unify_extracted
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s4 -> (projT1 args4)) ->
+ (projT1 args0) -> s10)%ptype
+ with
+ | Some (_, _, (_, _))%zrange =>
+ if
+ type.type_beq base.type
+ base.type.type_beq
+ ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
+ ((s4 -> (projT1 args4)) ->
+ (projT1 args0) -> s10)%ptype
+ then
+ v <- type.try_make_transport_cps s4 ℤ;
+ xv <- ident.unify
+ pattern.ident.Literal
+ ##(projT2 args4);
+ xv0 <- ident.unify
+ pattern.ident.Literal
+ ##(projT2 args0);
+ v0 <- type.try_make_transport_cps s10 ℤ;
+ fv <- (x12 <- (if
+ ((let (x12, _) :=
+ xv0 in
+ x12) =?
+ 2
+ ^ (2 *
+ (let (x12, _) :=
+ xv in
+ x12) / 2) - 1) &&
+ (ZRange.normalize
+ args6 >>
+ ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) :=
+ xv in
+ x12)) <=?
+ ZRange.normalize
+ args8)%zrange &&
+ (ZRange.normalize
+ (ZRange.constant
+ (let (x12, _) :=
+ xv0 in
+ x12)) &'
+ ZRange.normalize args <=?
+ ZRange.normalize
+ args3)%zrange &&
+ is_bounded_by_bool
+ (let (x12, _) :=
xv0 in
- x10)) &'
- ZRange.normalize args <=?
- ZRange.normalize args2)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_mulhl
- (2 *
- (let (x10, _) :=
- xv in
- x10)))%expr @
- (#(Z_cast args4)%expr @
- v
- (Compile.reflect x5),
- #(Z_cast args)%expr @
- v0
- (Compile.reflect x9))))%expr_pat
- else None);
- Some (Base x10));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
+ x12)
+ (ZRange.normalize
+ args1) &&
+ is_bounded_by_bool
+ (let (x12, _) :=
+ xv in
+ x12)
+ (ZRange.normalize
+ args5)
+ then
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_mulhl
+ (2 *
+ (let
+ (x12,
+ _) :=
+ xv in
+ x12)))%expr @
+ (#(Z_cast args6)%expr @
+ v
+ (Compile.reflect
+ x5),
+ #(Z_cast args)%expr @
+ v0
+ (Compile.reflect
+ x11))))%expr_pat
+ else None);
+ Some (Base x12));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | @expr.App _ _ _ s10 _ ($_)%expr _ | @expr.App
+ _ _ _ s10 _ (@expr.Abs _ _ _ _ _ _) _ |
+ @expr.App _ _ _ s10 _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s10 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
end
- | @expr.App _ _ _ s8 _ ($_)%expr _ | @expr.App _ _
- _ s8 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
- _ s8 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s8 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
- end
- | @expr.App _ _ _ s8 _ (@expr.Ident _ _ _ t6 idc6)
- x9 =>
- match x7 with
- | @expr.Ident _ _ _ t7 idc7 =>
- args <- invert_bind_args idc7 Raw.ident.Literal;
- args0 <- invert_bind_args idc6 Raw.ident.Z_cast;
- _ <- invert_bind_args idc5 Raw.ident.Z_land;
- args2 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc3
+ end;;
+ match x8 with
+ | (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9)%expr_pat =>
+ args <- invert_bind_args idc9 Raw.ident.Literal;
+ args0 <- invert_bind_args idc8 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc7 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc6 Raw.ident.Z_land;
+ args3 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc4
Raw.ident.Literal;
- args4 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args5 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args6 <- invert_bind_args idc2 Raw.ident.Z_cast;
_ <- invert_bind_args idc1 Raw.ident.Z_shiftr;
- args6 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ args8 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_mul;
match
pattern.type.unify_extracted
((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- ((s4 -> (projT1 args3)) ->
- s8 -> (projT1 args))%ptype
+ ((s4 -> (projT1 args4)) ->
+ s9 -> (projT1 args))%ptype
with
| Some (_, _, (_, _))%zrange =>
if
type.type_beq base.type base.type.type_beq
((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- ((s4 -> (projT1 args3)) ->
- s8 -> (projT1 args))%ptype
+ ((s4 -> (projT1 args4)) ->
+ s9 -> (projT1 args))%ptype
then
v <- type.try_make_transport_cps s4 ℤ;
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args3);
- v0 <- type.try_make_transport_cps s8 ℤ;
+ ##(projT2 args4);
+ v0 <- type.try_make_transport_cps s9 ℤ;
xv0 <- ident.unify pattern.ident.Literal
##(projT2 args);
- fv <- (x10 <- (if
- ((let (x10, _) := xv0 in
- x10) =?
+ fv <- (x12 <- (if
+ ((let (x12, _) := xv0 in
+ x12) =?
2
^ (2 *
- (let (x10, _) := xv in
- x10) / 2) - 1) &&
- (ZRange.normalize args4 >>
+ (let (x12, _) := xv in
+ x12) / 2) - 1) &&
+ (ZRange.normalize args6 >>
ZRange.normalize
(ZRange.constant
- (let (x10, _) :=
+ (let (x12, _) :=
xv in
- x10)) <=?
- ZRange.normalize args6)%zrange &&
- (ZRange.normalize args0 &'
+ x12)) <=?
+ ZRange.normalize args8)%zrange &&
+ (ZRange.normalize args1 &'
ZRange.normalize
(ZRange.constant
- (let (x10, _) :=
+ (let (x12, _) :=
xv0 in
- x10)) <=?
- ZRange.normalize args2)%zrange
+ x12)) <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x12, _) := xv0 in
+ x12)
+ (ZRange.normalize args0) &&
+ is_bounded_by_bool
+ (let (x12, _) := xv in
+ x12)
+ (ZRange.normalize args5)
then
Some
(#(Z_cast range)%expr @
(#(fancy_mulhl
(2 *
- (let (x10, _) :=
+ (let (x12, _) :=
xv in
- x10)))%expr @
- (#(Z_cast args4)%expr @
+ x12)))%expr @
+ (#(Z_cast args6)%expr @
v
(Compile.reflect x5),
- #(Z_cast args0)%expr @
+ #(Z_cast args1)%expr @
v0
- (Compile.reflect x9))))%expr_pat
+ (Compile.reflect x10))))%expr_pat
else None);
- Some (Base x10));
+ Some (Base x12));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
+ | (@expr.Ident _ _ _ t8 idc8 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Abs _ _ _ _ _
+ _)%expr_pat |
+ (@expr.Ident _ _ _ t8 idc8 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.LetIn _ _ _ _
+ _ _ _)%expr_pat => None
| _ => None
end
- | @expr.App _ _ _ s8 _ ($_)%expr _ | @expr.App _ _ _ s8
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s8 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s8 _
+ | @expr.App _ _ _ s9 _ ($_)%expr _ | @expr.App _ _ _ s9
+ _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s9 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s9 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end;;
- match x8 with
- | @expr.App _ _ _ s8 _ (@expr.Ident _ _ _ t6 idc6)
- x9 =>
- match x7 with
- | @expr.Ident _ _ _ t7 idc7 =>
- args <- invert_bind_args idc7 Raw.ident.Literal;
- args0 <- invert_bind_args idc6 Raw.ident.Z_cast;
- _ <- invert_bind_args idc5 Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc3
+ match x9 with
+ | @expr.App _ _ _ s9 _ (@expr.Ident _ _ _ t7 idc7)
+ x10 =>
+ match x8 with
+ | (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9)%expr_pat =>
+ args <- invert_bind_args idc9 Raw.ident.Literal;
+ args0 <- invert_bind_args idc8 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc7 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc6 Raw.ident.Z_shiftr;
+ args3 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc4
Raw.ident.Literal;
- args4 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args5 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args6 <- invert_bind_args idc2 Raw.ident.Z_cast;
_ <- invert_bind_args idc1 Raw.ident.Z_shiftr;
- args6 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ args8 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_mul;
match
pattern.type.unify_extracted
((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- ((s4 -> (projT1 args3)) ->
- s8 -> (projT1 args))%ptype
+ ((s4 -> (projT1 args4)) ->
+ s9 -> (projT1 args))%ptype
with
| Some (_, _, (_, _))%zrange =>
if
type.type_beq base.type base.type.type_beq
((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- ((s4 -> (projT1 args3)) ->
- s8 -> (projT1 args))%ptype
+ ((s4 -> (projT1 args4)) ->
+ s9 -> (projT1 args))%ptype
then
v <- type.try_make_transport_cps s4 ℤ;
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args3);
- v0 <- type.try_make_transport_cps s8 ℤ;
+ ##(projT2 args4);
+ v0 <- type.try_make_transport_cps s9 ℤ;
xv0 <- ident.unify pattern.ident.Literal
##(projT2 args);
- fv <- (x10 <- (if
- ((let (x10, _) := xv in
- x10) =?
- (let (x10, _) := xv0 in
- x10)) &&
- (ZRange.normalize args4 >>
+ fv <- (x12 <- (if
+ ((let (x12, _) := xv in
+ x12) =?
+ (let (x12, _) := xv0 in
+ x12)) &&
+ (ZRange.normalize args6 >>
ZRange.normalize
(ZRange.constant
- (let (x10, _) :=
+ (let (x12, _) :=
xv in
- x10)) <=?
- ZRange.normalize args6)%zrange &&
- (ZRange.normalize args0 >>
+ x12)) <=?
+ ZRange.normalize args8)%zrange &&
+ (ZRange.normalize args1 >>
ZRange.normalize
(ZRange.constant
- (let (x10, _) :=
+ (let (x12, _) :=
xv0 in
- x10)) <=?
- ZRange.normalize args2)%zrange
+ x12)) <=?
+ ZRange.normalize args3)%zrange &&
+ is_bounded_by_bool
+ (let (x12, _) := xv in
+ x12)
+ (ZRange.normalize args5) &&
+ is_bounded_by_bool
+ (let (x12, _) := xv0 in
+ x12)
+ (ZRange.normalize args0)
then
Some
(#(Z_cast range)%expr @
(#(fancy_mulhh
(2 *
- (let (x10, _) :=
+ (let (x12, _) :=
xv in
- x10)))%expr @
- (#(Z_cast args4)%expr @
+ x12)))%expr @
+ (#(Z_cast args6)%expr @
v
(Compile.reflect x5),
- #(Z_cast args0)%expr @
+ #(Z_cast args1)%expr @
v0
- (Compile.reflect x9))))%expr_pat
+ (Compile.reflect x10))))%expr_pat
else None);
- Some (Base x10));
+ Some (Base x12));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
+ | (@expr.Ident _ _ _ t8 idc8 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Abs _ _ _ _ _
+ _)%expr_pat |
+ (@expr.Ident _ _ _ t8 idc8 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.LetIn _ _ _ _
+ _ _ _)%expr_pat => None
| _ => None
end
- | @expr.App _ _ _ s8 _ ($_)%expr _ | @expr.App _ _ _ s8
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s8 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s8 _
+ | @expr.App _ _ _ s9 _ ($_)%expr _ | @expr.App _ _ _ s9
+ _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s9 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s9 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
- | (@expr.Ident _ _ _ t4 idc4 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
+ | (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
None
| _ => None
end
+ | (@expr.Ident _ _ _ t3 idc3 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end
| @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4 _
@@ -1788,30 +2537,6 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
end
| @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0)) #(_)%expr_pat | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0)) ($_)%expr | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0)) (@expr.Abs _ _ _ _ _ _) | @expr.App _ _
- _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0)) (($_)%expr @ _)%expr_pat | @expr.App _
- _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0)) (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat |
- @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0)) (_ @ _ @ _)%expr_pat | @expr.App _ _ _
- s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0))
- (@expr.LetIn _ _ _ _ _ _ _ @ _)%expr_pat | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0)) (@expr.LetIn _ _ _ _ _ _ _) | @expr.App
- _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
- (@expr.Ident _ _ _ t0 idc0 @ #(_))%expr_pat) _ | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
(@expr.Ident _ _ _ t0 idc0 @ ($_)%expr)%expr_pat) _ | @expr.App _ _
_ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
@@ -1845,6 +2570,8 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(@expr.Ident _ _ _ t0 idc0 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat)
_ => None
| @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc) #(_)%expr_pat) _ |
+ @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc) ($_)%expr) _ |
@expr.App _ _ _ s _
(@expr.App _ _ _ s0 _ (@expr.Ident _ _ _ t idc)
@@ -1866,524 +2593,202 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
| (@expr.Ident _ _ _ t0 idc0 @
(@expr.Ident _ _ _ t1 idc1 @ x5 @ x4))%expr_pat =>
match x5 with
- | @expr.Ident _ _ _ t2 idc2 =>
+ | (@expr.Ident _ _ _ t2 idc2 @ @expr.Ident _ _ _ t3 idc3)%expr_pat =>
match x4 with
- | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6 =>
+ | @expr.App _ _ _ s6 _ (@expr.Ident _ _ _ t4 idc4) x7 =>
match x1 with
- | @expr.Ident _ _ _ t4 idc4 =>
+ | @expr.App _ _ _ s7 _ (@expr.Ident _ _ _ t5 idc5)
+ x8 =>
match x0 with
- | @expr.Ident _ _ _ t5 idc5 =>
- args <- invert_bind_args idc5 Raw.ident.Literal;
- args0 <- invert_bind_args idc4
- Raw.ident.Literal;
- args1 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc2
- Raw.ident.Literal;
- _ <- invert_bind_args idc1 Raw.ident.Z_cc_m;
- args4 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_zselect;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) -> (projT1 args0)) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) ->
- (projT1 args0)) -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s5 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x7 <- (if
- ((let (x7, _) := xv in x7) =?
- 2
- ^ Z.log2
- (let (x7, _) := xv in
- x7)) &&
- ((ZRange.cc_m
- (let (x7, _) := xv in
- x7))
- (ZRange.normalize args1) <=?
- ZRange.normalize args4)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_selm
- (Z.log2
- (let (x7, _) :=
- xv in
- x7)))%expr @
- (#(Z_cast args1)%expr @
- v (Compile.reflect x6),
- (##(let (x7, _) :=
- xv0 in
- x7))%expr,
- (##(let (x7, _) :=
- xv1 in
- x7))%expr)))%expr_pat
- else None);
- Some (Base x7));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s6 _ (@expr.Ident _ _ _ t5 idc5)
- x7 =>
- args <- invert_bind_args idc5 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc4
- Raw.ident.Literal;
- args1 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc2
- Raw.ident.Literal;
- _ <- invert_bind_args idc1 Raw.ident.Z_cc_m;
- args4 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_zselect;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) -> (projT1 args0)) ->
- s6)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) ->
- (projT1 args0)) -> s6)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s5 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v0 <- type.try_make_transport_cps s6 ℤ;
- fv <- (x8 <- (if
- ((let (x8, _) := xv in x8) =?
- 2
- ^ Z.log2
- (let (x8, _) := xv in
- x8)) &&
- ((ZRange.cc_m
- (let (x8, _) := xv in
- x8))
- (ZRange.normalize args1) <=?
- ZRange.normalize args4)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_selm
- (Z.log2
- (let (x8, _) :=
- xv in
- x8)))%expr @
- (#(Z_cast args1)%expr @
- v (Compile.reflect x6),
- (##(let (x8, _) :=
- xv0 in
- x8))%expr,
- #(Z_cast args)%expr @
- v0 (Compile.reflect x7))))%expr_pat
- else None);
- Some (Base x8));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s6 _ ($_)%expr _ | @expr.App _ _
- _ s6 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
- _ s6 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s6 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s6 _ (@expr.Ident _ _ _ t4 idc4)
- x7 =>
- match x0 with
- | @expr.Ident _ _ _ t5 idc5 =>
- args <- invert_bind_args idc5 Raw.ident.Literal;
- args0 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc2
- Raw.ident.Literal;
- _ <- invert_bind_args idc1 Raw.ident.Z_cc_m;
- args4 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_zselect;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) -> s6) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) -> s6) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s5 ℤ;
- v0 <- type.try_make_transport_cps s6 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x8 <- (if
- ((let (x8, _) := xv in x8) =?
- 2
- ^ Z.log2
- (let (x8, _) := xv in
- x8)) &&
- ((ZRange.cc_m
- (let (x8, _) := xv in
- x8))
- (ZRange.normalize args1) <=?
- ZRange.normalize args4)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_selm
- (Z.log2
- (let (x8, _) :=
- xv in
- x8)))%expr @
- (#(Z_cast args1)%expr @
- v (Compile.reflect x6),
- #(Z_cast args0)%expr @
- v0 (Compile.reflect x7),
- (##(let (x8, _) :=
- xv0 in
- x8))%expr)))%expr_pat
- else None);
- Some (Base x8));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s7 _ (@expr.Ident _ _ _ t5 idc5)
- x8 =>
- args <- invert_bind_args idc5 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc2
+ | @expr.App _ _ _ s8 _ (@expr.Ident _ _ _ t6 idc6)
+ x9 =>
+ args <- invert_bind_args idc6 Raw.ident.Z_cast;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc3
Raw.ident.Literal;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
_ <- invert_bind_args idc1 Raw.ident.Z_cc_m;
- args4 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_zselect;
match
pattern.type.unify_extracted
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) -> s6) -> s7)%ptype
+ ((((projT1 args2) -> s6) -> s7) -> s8)%ptype
with
| Some (_, _, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) -> s6) -> s7)%ptype
+ ((((projT1 args2) -> s6) -> s7) -> s8)%ptype
then
xv <- ident.unify pattern.ident.Literal
##(projT2 args2);
- v <- type.try_make_transport_cps s5 ℤ;
- v0 <- type.try_make_transport_cps s6 ℤ;
- v1 <- type.try_make_transport_cps s7 ℤ;
- fv <- (x9 <- (if
- ((let (x9, _) := xv in x9) =?
- 2
- ^ Z.log2
- (let (x9, _) := xv in
- x9)) &&
- ((ZRange.cc_m
- (let (x9, _) := xv in
- x9))
- (ZRange.normalize args1) <=?
- ZRange.normalize args4)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_selm
- (Z.log2
- (let (x9, _) :=
- xv in
- x9)))%expr @
- (#(Z_cast args1)%expr @
- v (Compile.reflect x6),
- #(Z_cast args0)%expr @
- v0 (Compile.reflect x7),
- #(Z_cast args)%expr @
- v1 (Compile.reflect x8))))%expr_pat
- else None);
- Some (Base x9));
+ v <- type.try_make_transport_cps s6 ℤ;
+ v0 <- type.try_make_transport_cps s7 ℤ;
+ v1 <- type.try_make_transport_cps s8 ℤ;
+ fv <- (x10 <- (if
+ ((let (x10, _) := xv in
+ x10) =?
+ 2
+ ^ Z.log2
+ (let (x10, _) :=
+ xv in
+ x10)) &&
+ ((ZRange.cc_m
+ (let (x10, _) := xv in
+ x10))
+ (ZRange.normalize args1) <=?
+ ZRange.normalize args5)%zrange &&
+ is_bounded_by_bool
+ (let (x10, _) := xv in
+ x10)
+ (ZRange.normalize args3)
+ then
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_selm
+ (Z.log2
+ (let
+ (x10, _) :=
+ xv in
+ x10)))%expr @
+ (#(Z_cast args1)%expr @
+ v
+ (Compile.reflect x7),
+ #(Z_cast args0)%expr @
+ v0
+ (Compile.reflect x8),
+ #(Z_cast args)%expr @
+ v1
+ (Compile.reflect x9))))%expr_pat
+ else None);
+ Some (Base x10));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | @expr.App _ _ _ s7 _ ($_)%expr _ | @expr.App _ _
- _ s7 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
- _ s7 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s7 _
+ | @expr.App _ _ _ s8 _ ($_)%expr _ | @expr.App _ _
+ _ s8 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
+ _ s8 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s8 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
- | @expr.App _ _ _ s6 _ ($_)%expr _ | @expr.App _ _ _ s6
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s6 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s6 _
+ | @expr.App _ _ _ s7 _ ($_)%expr _ | @expr.App _ _ _ s7
+ _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s7 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s7 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
- | @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _ _ s5 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s5 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s5 _
+ | @expr.App _ _ _ s6 _ ($_)%expr _ | @expr.App _ _ _ s6 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s6 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s6 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
+ | (@expr.Ident _ _ _ t2 idc2 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t2 idc2 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end;;
match x5 with
- | @expr.Ident _ _ _ t2 idc2 =>
- match x4 with
- | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6 =>
- match x1 with
- | @expr.Ident _ _ _ t4 idc4 =>
- match x0 with
- | @expr.Ident _ _ _ t5 idc5 =>
- args <- invert_bind_args idc5 Raw.ident.Literal;
- args0 <- invert_bind_args idc4
- Raw.ident.Literal;
- args1 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc2
- Raw.ident.Literal;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args4 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_zselect;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) -> (projT1 args0)) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) ->
- (projT1 args0)) -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s5 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x7 <- (if
- ((let (x7, _) := xv in x7) =?
- 1) &&
- (ZRange.normalize
- (ZRange.constant
- (let (x7, _) := xv in
- x7)) &'
- ZRange.normalize args1 <=?
- ZRange.normalize args4)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_sell)%expr @
- (#(Z_cast args1)%expr @
- v (Compile.reflect x6),
- (##(let (x7, _) :=
- xv0 in
- x7))%expr,
- (##(let (x7, _) :=
- xv1 in
- x7))%expr)))%expr_pat
- else None);
- Some (Base x7));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s6 _ (@expr.Ident _ _ _ t5 idc5)
- x7 =>
- args <- invert_bind_args idc5 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc4
- Raw.ident.Literal;
- args1 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc2
- Raw.ident.Literal;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args4 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_zselect;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) -> (projT1 args0)) ->
- s6)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) ->
- (projT1 args0)) -> s6)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s5 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v0 <- type.try_make_transport_cps s6 ℤ;
- fv <- (x8 <- (if
- ((let (x8, _) := xv in x8) =?
- 1) &&
- (ZRange.normalize
- (ZRange.constant
- (let (x8, _) := xv in
- x8)) &'
- ZRange.normalize args1 <=?
- ZRange.normalize args4)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_sell)%expr @
- (#(Z_cast args1)%expr @
- v (Compile.reflect x6),
- (##(let (x8, _) :=
- xv0 in
- x8))%expr,
- #(Z_cast args)%expr @
- v0 (Compile.reflect x7))))%expr_pat
- else None);
- Some (Base x8));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s6 _ ($_)%expr _ | @expr.App _ _
- _ s6 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
- _ s6 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s6 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
+ | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t2 idc2) x6 =>
+ match x6 with
+ | @expr.Ident _ _ _ t3 idc3 =>
+ match x4 with
| @expr.App _ _ _ s6 _ (@expr.Ident _ _ _ t4 idc4)
x7 =>
- match x0 with
- | @expr.Ident _ _ _ t5 idc5 =>
- args <- invert_bind_args idc5 Raw.ident.Literal;
- args0 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc2
- Raw.ident.Literal;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args4 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_zselect;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) -> s6) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) -> s6) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s5 ℤ;
- v0 <- type.try_make_transport_cps s6 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x8 <- (if
- ((let (x8, _) := xv in x8) =?
- 1) &&
- (ZRange.normalize
- (ZRange.constant
- (let (x8, _) := xv in
- x8)) &'
- ZRange.normalize args1 <=?
- ZRange.normalize args4)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_sell)%expr @
- (#(Z_cast args1)%expr @
- v (Compile.reflect x6),
- #(Z_cast args0)%expr @
- v0 (Compile.reflect x7),
- (##(let (x8, _) :=
- xv0 in
- x8))%expr)))%expr_pat
- else None);
- Some (Base x8));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
+ match x1 with
| @expr.App _ _ _ s7 _ (@expr.Ident _ _ _ t5 idc5)
x8 =>
- args <- invert_bind_args idc5 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc2
- Raw.ident.Literal;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args4 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_zselect;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) -> s6) -> s7)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s5) -> s6) -> s7)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s5 ℤ;
- v0 <- type.try_make_transport_cps s6 ℤ;
- v1 <- type.try_make_transport_cps s7 ℤ;
- fv <- (x9 <- (if
- ((let (x9, _) := xv in x9) =?
- 1) &&
- (ZRange.normalize
- (ZRange.constant
- (let (x9, _) := xv in
- x9)) &'
- ZRange.normalize args1 <=?
- ZRange.normalize args4)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_sell)%expr @
- (#(Z_cast args1)%expr @
- v (Compile.reflect x6),
- #(Z_cast args0)%expr @
- v0 (Compile.reflect x7),
- #(Z_cast args)%expr @
- v1 (Compile.reflect x8))))%expr_pat
- else None);
- Some (Base x9));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
+ match x0 with
+ | @expr.App _ _ _ s8 _
+ (@expr.Ident _ _ _ t6 idc6) x9 =>
+ args <- invert_bind_args idc6
+ Raw.ident.Z_cast;
+ args0 <- invert_bind_args idc5
+ Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4
+ Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc3
+ Raw.ident.Literal;
+ args3 <- invert_bind_args idc2
+ Raw.ident.Z_cast;
+ _ <- invert_bind_args idc1 Raw.ident.Z_land;
+ args5 <- invert_bind_args idc0
+ Raw.ident.Z_cast;
+ _ <- invert_bind_args idc
+ Raw.ident.Z_zselect;
+ match
+ pattern.type.unify_extracted
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args2) -> s6) -> s7) -> s8)%ptype
+ with
+ | Some (_, _, _, _)%zrange =>
+ if
+ type.type_beq base.type
+ base.type.type_beq
+ (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
+ ((((projT1 args2) -> s6) -> s7) ->
+ s8)%ptype
+ then
+ xv <- ident.unify
+ pattern.ident.Literal
+ ##(projT2 args2);
+ v <- type.try_make_transport_cps s6 ℤ;
+ v0 <- type.try_make_transport_cps s7 ℤ;
+ v1 <- type.try_make_transport_cps s8 ℤ;
+ fv <- (x10 <- (if
+ ((let (x10, _) :=
+ xv in
+ x10) =? 1) &&
+ (ZRange.normalize
+ (ZRange.constant
+ (let (x10, _) :=
+ xv in
+ x10)) &'
+ ZRange.normalize
+ args1 <=?
+ ZRange.normalize
+ args5)%zrange &&
+ is_bounded_by_bool
+ (let (x10, _) :=
+ xv in
+ x10)
+ (ZRange.normalize
+ args3)
+ then
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_sell)%expr @
+ (#(Z_cast args1)%expr @
+ v
+ (Compile.reflect
+ x7),
+ #(Z_cast args0)%expr @
+ v0
+ (Compile.reflect
+ x8),
+ #(Z_cast args)%expr @
+ v1
+ (Compile.reflect
+ x9))))%expr_pat
+ else None);
+ Some (Base x10));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | @expr.App _ _ _ s8 _ ($_)%expr _ | @expr.App
+ _ _ _ s8 _ (@expr.Abs _ _ _ _ _ _) _ |
+ @expr.App _ _ _ s8 _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s8 _
+ (@expr.LetIn _ _ _ _ _ _ _) _ => None
+ | _ => None
end
| @expr.App _ _ _ s7 _ ($_)%expr _ | @expr.App _ _
_ s7 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
@@ -2397,261 +2802,92 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
- | @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _ _ s5 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s5 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s5 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
- end
- | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t2 idc2) x6 =>
+ end;;
match x4 with
- | @expr.Ident _ _ _ t3 idc3 =>
+ | (@expr.Ident _ _ _ t3 idc3 @ @expr.Ident _ _ _ t4 idc4)%expr_pat =>
match x1 with
- | @expr.Ident _ _ _ t4 idc4 =>
+ | @expr.App _ _ _ s7 _ (@expr.Ident _ _ _ t5 idc5)
+ x8 =>
match x0 with
- | @expr.Ident _ _ _ t5 idc5 =>
- args <- invert_bind_args idc5 Raw.ident.Literal;
- args0 <- invert_bind_args idc4
- Raw.ident.Literal;
- args1 <- invert_bind_args idc3
- Raw.ident.Literal;
- args2 <- invert_bind_args idc2 Raw.ident.Z_cast;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args4 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_zselect;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s5 -> (projT1 args1)) -> (projT1 args0)) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s5 -> (projT1 args1)) ->
- (projT1 args0)) -> (projT1 args))%ptype
- then
- v <- type.try_make_transport_cps s5 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x7 <- (if
- ((let (x7, _) := xv in x7) =?
- 1) &&
- (ZRange.normalize args2 &'
- ZRange.normalize
- (ZRange.constant
- (let (x7, _) := xv in
- x7)) <=?
- ZRange.normalize args4)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_sell)%expr @
- (#(Z_cast args2)%expr @
- v (Compile.reflect x6),
- (##(let (x7, _) :=
- xv0 in
- x7))%expr,
- (##(let (x7, _) :=
- xv1 in
- x7))%expr)))%expr_pat
- else None);
- Some (Base x7));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s6 _ (@expr.Ident _ _ _ t5 idc5)
- x7 =>
- args <- invert_bind_args idc5 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc4
- Raw.ident.Literal;
- args1 <- invert_bind_args idc3
- Raw.ident.Literal;
- args2 <- invert_bind_args idc2 Raw.ident.Z_cast;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args4 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_zselect;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s5 -> (projT1 args1)) -> (projT1 args0)) ->
- s6)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s5 -> (projT1 args1)) ->
- (projT1 args0)) -> s6)%ptype
- then
- v <- type.try_make_transport_cps s5 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v0 <- type.try_make_transport_cps s6 ℤ;
- fv <- (x8 <- (if
- ((let (x8, _) := xv in x8) =?
- 1) &&
- (ZRange.normalize args2 &'
- ZRange.normalize
- (ZRange.constant
- (let (x8, _) := xv in
- x8)) <=?
- ZRange.normalize args4)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_sell)%expr @
- (#(Z_cast args2)%expr @
- v (Compile.reflect x6),
- (##(let (x8, _) :=
- xv0 in
- x8))%expr,
- #(Z_cast args)%expr @
- v0 (Compile.reflect x7))))%expr_pat
- else None);
- Some (Base x8));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s6 _ ($_)%expr _ | @expr.App _ _
- _ s6 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
- _ s6 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s6 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s6 _ (@expr.Ident _ _ _ t4 idc4)
- x7 =>
- match x0 with
- | @expr.Ident _ _ _ t5 idc5 =>
- args <- invert_bind_args idc5 Raw.ident.Literal;
- args0 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc3
- Raw.ident.Literal;
- args2 <- invert_bind_args idc2 Raw.ident.Z_cast;
- _ <- invert_bind_args idc1 Raw.ident.Z_land;
- args4 <- invert_bind_args idc0 Raw.ident.Z_cast;
- _ <- invert_bind_args idc Raw.ident.Z_zselect;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s5 -> (projT1 args1)) -> s6) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s5 -> (projT1 args1)) -> s6) ->
- (projT1 args))%ptype
- then
- v <- type.try_make_transport_cps s5 ℤ;
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v0 <- type.try_make_transport_cps s6 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x8 <- (if
- ((let (x8, _) := xv in x8) =?
- 1) &&
- (ZRange.normalize args2 &'
- ZRange.normalize
- (ZRange.constant
- (let (x8, _) := xv in
- x8)) <=?
- ZRange.normalize args4)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_sell)%expr @
- (#(Z_cast args2)%expr @
- v (Compile.reflect x6),
- #(Z_cast args0)%expr @
- v0 (Compile.reflect x7),
- (##(let (x8, _) :=
- xv0 in
- x8))%expr)))%expr_pat
- else None);
- Some (Base x8));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s7 _ (@expr.Ident _ _ _ t5 idc5)
- x8 =>
- args <- invert_bind_args idc5 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc4 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc3
+ | @expr.App _ _ _ s8 _ (@expr.Ident _ _ _ t6 idc6)
+ x9 =>
+ args <- invert_bind_args idc6 Raw.ident.Z_cast;
+ args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc4
Raw.ident.Literal;
- args2 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
_ <- invert_bind_args idc1 Raw.ident.Z_land;
- args4 <- invert_bind_args idc0 Raw.ident.Z_cast;
+ args5 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_zselect;
match
pattern.type.unify_extracted
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s5 -> (projT1 args1)) -> s6) -> s7)%ptype
+ (((s5 -> (projT1 args1)) -> s7) -> s8)%ptype
with
| Some (_, _, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((s5 -> (projT1 args1)) -> s6) -> s7)%ptype
+ (((s5 -> (projT1 args1)) -> s7) -> s8)%ptype
then
v <- type.try_make_transport_cps s5 ℤ;
xv <- ident.unify pattern.ident.Literal
##(projT2 args1);
- v0 <- type.try_make_transport_cps s6 ℤ;
- v1 <- type.try_make_transport_cps s7 ℤ;
- fv <- (x9 <- (if
- ((let (x9, _) := xv in x9) =?
- 1) &&
- (ZRange.normalize args2 &'
- ZRange.normalize
- (ZRange.constant
- (let (x9, _) := xv in
- x9)) <=?
- ZRange.normalize args4)%zrange
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_sell)%expr @
- (#(Z_cast args2)%expr @
- v (Compile.reflect x6),
- #(Z_cast args0)%expr @
- v0 (Compile.reflect x7),
- #(Z_cast args)%expr @
- v1 (Compile.reflect x8))))%expr_pat
- else None);
- Some (Base x9));
+ v0 <- type.try_make_transport_cps s7 ℤ;
+ v1 <- type.try_make_transport_cps s8 ℤ;
+ fv <- (x10 <- (if
+ ((let (x10, _) := xv in
+ x10) =? 1) &&
+ (ZRange.normalize args3 &'
+ ZRange.normalize
+ (ZRange.constant
+ (let (x10, _) :=
+ xv in
+ x10)) <=?
+ ZRange.normalize args5)%zrange &&
+ is_bounded_by_bool
+ (let (x10, _) := xv in
+ x10)
+ (ZRange.normalize args2)
+ then
+ Some
+ (#(Z_cast range)%expr @
+ (#(fancy_sell)%expr @
+ (#(Z_cast args3)%expr @
+ v
+ (Compile.reflect x6),
+ #(Z_cast args0)%expr @
+ v0
+ (Compile.reflect x8),
+ #(Z_cast args)%expr @
+ v1
+ (Compile.reflect x9))))%expr_pat
+ else None);
+ Some (Base x10));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | @expr.App _ _ _ s7 _ ($_)%expr _ | @expr.App _ _
- _ s7 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
- _ s7 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s7 _
+ | @expr.App _ _ _ s8 _ ($_)%expr _ | @expr.App _ _
+ _ s8 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
+ _ s8 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s8 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
- | @expr.App _ _ _ s6 _ ($_)%expr _ | @expr.App _ _ _ s6
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s6 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s6 _
+ | @expr.App _ _ _ s7 _ ($_)%expr _ | @expr.App _ _ _ s7
+ _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s7 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s7 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
+ | (@expr.Ident _ _ _ t3 idc3 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t3 idc3 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end
| @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _ _ s5 _
@@ -2700,43 +2936,51 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
| @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.Ident _ _ _ t1 idc1))
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t2 idc2) x4))
- (@expr.Ident _ _ _ t3 idc3) =>
- args <- invert_bind_args idc3 Raw.ident.Literal;
- args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc1 Raw.ident.Literal;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6))
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.Ident _ _ _ t5 idc5)%expr_pat =>
+ args <- invert_bind_args idc5 Raw.ident.Literal;
+ args0 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc1 Raw.ident.Literal;
+ args4 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_rshi;
match
pattern.type.unify_extracted (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) -> s3) -> (projT1 args))%ptype
+ ((((projT1 args3) -> s4) -> s5) -> (projT1 args))%ptype
with
| Some (_, _, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) -> s3) -> (projT1 args))%ptype
+ ((((projT1 args3) -> s4) -> s5) -> (projT1 args))%ptype
then
- xv <- ident.unify pattern.ident.Literal ##(projT2 args2);
- xv0 <- ident.unify pattern.ident.Literal ##(projT2 args1);
- v <- type.try_make_transport_cps s3 ℤ;
- xv1 <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (x5 <- (if
- (let (x5, _) := xv in x5) =?
- 2 ^ Z.log2 (let (x5, _) := xv in x5)
+ xv <- ident.unify pattern.ident.Literal ##(projT2 args3);
+ v <- type.try_make_transport_cps s4 ℤ;
+ v0 <- type.try_make_transport_cps s5 ℤ;
+ xv0 <- ident.unify pattern.ident.Literal ##(projT2 args);
+ fv <- (x8 <- (if
+ ((let (x8, _) := xv in x8) =?
+ 2 ^ Z.log2 (let (x8, _) := xv in x8)) &&
+ is_bounded_by_bool (let (x8, _) := xv in x8)
+ (ZRange.normalize args4) &&
+ is_bounded_by_bool (let (x8, _) := xv0 in x8)
+ (ZRange.normalize args0)
then
Some
(#(Z_cast range)%expr @
(#(fancy_rshi
- (Z.log2 (let (x5, _) := xv in x5))
- (let (x5, _) := xv1 in x5))%expr @
- ((##(let (x5, _) := xv0 in x5))%expr,
- #(Z_cast args0)%expr @
- v (Compile.reflect x4))))%expr_pat
+ (Z.log2 (let (x8, _) := xv in x8))
+ (let (x8, _) := xv0 in x8))%expr @
+ (#(Z_cast args2)%expr @
+ v (Compile.reflect x5),
+ #(Z_cast args1)%expr @
+ v0 (Compile.reflect x6))))%expr_pat
else None);
- Some (Base x5));
+ Some (Base x8));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
@@ -2745,276 +2989,211 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
| @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.Ident _ _ _ t1 idc1)) (@expr.App _ _ _ s3 _ ($_)%expr _))
- _ | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _
- (@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.Ident _ _ _ t1 idc1))
- (@expr.App _ _ _ s3 _ (@expr.Abs _ _ _ _ _ _) _)) _ | @expr.App _ _
- _ s _
- (@expr.App _ _ _ s0 _
- (@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.Ident _ _ _ t1 idc1))
- (@expr.App _ _ _ s3 _ (_ @ _)%expr_pat _)) _ | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _
- (@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.Ident _ _ _ t1 idc1))
- (@expr.App _ _ _ s3 _ (@expr.LetIn _ _ _ _ _ _ _) _)) _ => None
- | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _
- (@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.Ident _ _ _ t1 idc1)) #(_)%expr_pat) _ | @expr.App _ _ _ s _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6))
+ (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr)%expr_pat | @expr.App _ _ _ s
+ _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.Ident _ _ _ t1 idc1)) ($_)%expr) _ | @expr.App _ _ _ s _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6))
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.Ident _ _ _ t1 idc1)) (@expr.Abs _ _ _ _ _ _)) _ | @expr.App
- _ _ _ s _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6))
+ (@expr.Ident _ _ _ t4 idc4 @ (_ @ _))%expr_pat | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.Ident _ _ _ t1 idc1)) (@expr.LetIn _ _ _ _ _ _ _)) _ => None
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6))
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.Ident _ _ _ t2 idc2)) (@expr.Ident _ _ _ t3 idc3) =>
- args <- invert_bind_args idc3 Raw.ident.Literal;
- args0 <- invert_bind_args idc2 Raw.ident.Literal;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_rshi;
- match
- pattern.type.unify_extracted (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> (projT1 args0)) -> (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> (projT1 args0)) -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal ##(projT2 args2);
- v <- type.try_make_transport_cps s3 ℤ;
- xv0 <- ident.unify pattern.ident.Literal ##(projT2 args0);
- xv1 <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (x5 <- (if
- (let (x5, _) := xv in x5) =?
- 2 ^ Z.log2 (let (x5, _) := xv in x5)
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_rshi
- (Z.log2 (let (x5, _) := xv in x5))
- (let (x5, _) := xv1 in x5))%expr @
- (#(Z_cast args1)%expr @
- v (Compile.reflect x4),
- (##(let (x5, _) := xv0 in x5))%expr)))%expr_pat
- else None);
- Some (Base x5));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6))
+ #(_)%expr_pat | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.Ident _ _ _ t1 idc1))
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t2 idc2) x4)) ($_)%expr |
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6)) ($_)%expr |
@expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.Ident _ _ _ t1 idc1))
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t2 idc2) x4))
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6))
(@expr.Abs _ _ _ _ _ _) | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.Ident _ _ _ t1 idc1))
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t2 idc2) x4))
- (_ @ _)%expr_pat | @expr.App _ _ _ s _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6))
+ (($_)%expr @ _)%expr_pat | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.Ident _ _ _ t1 idc1))
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t2 idc2) x4))
- (@expr.LetIn _ _ _ _ _ _ _) | @expr.App _ _ _ s _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6))
+ (@expr.Abs _ _ _ _ _ _ @ _)%expr_pat | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.Ident _ _ _ t2 idc2)) ($_)%expr | @expr.App _ _ _ s _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6))
+ (_ @ _ @ _)%expr_pat | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.Ident _ _ _ t2 idc2)) (@expr.Abs _ _ _ _ _ _) | @expr.App _ _
- _ s _
- (@expr.App _ _ _ s0 _
- (@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.Ident _ _ _ t2 idc2)) (_ @ _)%expr_pat | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _
- (@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.Ident _ _ _ t2 idc2)) (@expr.LetIn _ _ _ _ _ _ _) => None
- | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _
- (@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
- (@expr.Ident _ _ _ t3 idc3) =>
- args <- invert_bind_args idc3 Raw.ident.Literal;
- args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_rshi;
- match
- pattern.type.unify_extracted (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> s4) -> (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> s4) -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal ##(projT2 args2);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s4 ℤ;
- xv0 <- ident.unify pattern.ident.Literal ##(projT2 args);
- fv <- (x6 <- (if
- (let (x6, _) := xv in x6) =?
- 2 ^ Z.log2 (let (x6, _) := xv in x6)
- then
- Some
- (#(Z_cast range)%expr @
- (#(fancy_rshi
- (Z.log2 (let (x6, _) := xv in x6))
- (let (x6, _) := xv0 in x6))%expr @
- (#(Z_cast args1)%expr @
- v (Compile.reflect x4),
- #(Z_cast args0)%expr @
- v0 (Compile.reflect x5))))%expr_pat
- else None);
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _
- (@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5)) ($_)%expr |
- @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _
- (@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
- (@expr.Abs _ _ _ _ _ _) | @expr.App _ _ _ s _
- (@expr.App _ _ _ s0 _
- (@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
- (_ @ _)%expr_pat | @expr.App _ _ _ s _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6))
+ (@expr.LetIn _ _ _ _ _ _ _ @ _)%expr_pat | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6))
(@expr.LetIn _ _ _ _ _ _ _) => None
| @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.App _ _ _ s4 _ ($_)%expr _)) _ | @expr.App _ _ _ s _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ ($_)%expr _)) _ | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.App _ _ _ s4 _ (@expr.Abs _ _ _ _ _ _) _)) _ | @expr.App _ _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.Abs _ _ _ _ _ _) _)) _ | @expr.App _ _
_ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.App _ _ _ s4 _ (_ @ _)%expr_pat _)) _ | @expr.App _ _ _ s _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (_ @ _)%expr_pat _)) _ | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
- (@expr.App _ _ _ s4 _ (@expr.LetIn _ _ _ _ _ _ _) _)) _ => None
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ (@expr.App _ _ _ s5 _ (@expr.LetIn _ _ _ _ _ _ _) _)) _ => None
| @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4)) ($_)%expr)
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
+ #(_)%expr_pat) _ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _
+ (@expr.App _ _ _ s1 _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5)) ($_)%expr)
_ | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
(@expr.Abs _ _ _ _ _ _)) _ | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4))
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5))
(@expr.LetIn _ _ _ _ _ _ _)) _ => None
| @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ ($_)%expr _)) _) _ | @expr.App _ _ _ s _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ ($_)%expr _)) _) _ | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.Abs _ _ _ _ _ _) _)) _) _ | @expr.App
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.Abs _ _ _ _ _ _) _)) _) _ | @expr.App
_ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (_ @ _)%expr_pat _)) _) _ | @expr.App _ _ _ s
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (_ @ _)%expr_pat _)) _) _ | @expr.App _ _ _ s
_
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
- (@expr.App _ _ _ s3 _ (@expr.LetIn _ _ _ _ _ _ _) _)) _) _ => None
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ (@expr.App _ _ _ s4 _ (@expr.LetIn _ _ _ _ _ _ _) _)) _) _ => None
| @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
+ #(_)%expr_pat) _) _ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _
+ (@expr.App _ _ _ s1 _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
($_)%expr) _) _ | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
(@expr.Abs _ _ _ _ _ _)) _) _ | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
- (@expr.Ident _ _ _ t idc @ @expr.Ident _ _ _ t0 idc0)%expr_pat
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1))%expr_pat
(@expr.LetIn _ _ _ _ _ _ _)) _) _ => None
| @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
+ (@expr.App _ _ _ s1 _
+ (@expr.Ident _ _ _ t idc @ (@expr.Ident _ _ _ t0 idc0 @ ($_)%expr))%expr_pat
+ _) _) _ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _
+ (@expr.App _ _ _ s1 _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Abs _ _ _ _ _ _))%expr_pat _)
+ _) _ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _
+ (@expr.App _ _ _ s1 _
+ (@expr.Ident _ _ _ t idc @ (@expr.Ident _ _ _ t0 idc0 @ (_ @ _)))%expr_pat
+ _) _) _ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _
+ (@expr.App _ _ _ s1 _
+ (@expr.Ident _ _ _ t idc @
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.LetIn _ _ _ _ _ _ _))%expr_pat
+ _) _) _ => None
+ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _
+ (@expr.App _ _ _ s1 _ (@expr.Ident _ _ _ t idc @ #(_))%expr_pat _)
+ _) _ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _ (@expr.Ident _ _ _ t idc @ ($_)%expr)%expr_pat
_) _) _ | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
@@ -3022,7 +3201,20 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(@expr.Ident _ _ _ t idc @ @expr.Abs _ _ _ _ _ _)%expr_pat _) _)
_ | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
- (@expr.App _ _ _ s1 _ (@expr.Ident _ _ _ t idc @ (_ @ _))%expr_pat
+ (@expr.App _ _ _ s1 _
+ (@expr.Ident _ _ _ t idc @ (($_)%expr @ _))%expr_pat _) _) _ |
+ @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _
+ (@expr.App _ _ _ s1 _
+ (@expr.Ident _ _ _ t idc @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat _)
+ _) _ | @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _
+ (@expr.App _ _ _ s1 _
+ (@expr.Ident _ _ _ t idc @ (_ @ _ @ _))%expr_pat _) _) _ |
+ @expr.App _ _ _ s _
+ (@expr.App _ _ _ s0 _
+ (@expr.App _ _ _ s1 _
+ (@expr.Ident _ _ _ t idc @ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat
_) _) _ | @expr.App _ _ _ s _
(@expr.App _ _ _ s0 _
(@expr.App _ _ _ s1 _
@@ -3077,700 +3269,65 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
((match x with
| (@expr.Ident _ _ _ t idc @ x2 @ x1 @ x0)%expr_pat =>
match x2 with
- | @expr.Ident _ _ _ t0 idc0 =>
+ | (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat =>
match x1 with
- | @expr.Ident _ _ _ t1 idc1 =>
- match x0 with
- | @expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t2 idc2) x3 =>
- match x3 with
- | (@expr.Ident _ _ _ t3 idc3 @ x5 @ x4)%expr_pat =>
- match x5 with
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ @expr.Ident
- _ _ _ t7 idc7))%expr_pat =>
- match x4 with
- | @expr.Ident _ _ _ t8 idc8 =>
- args <- invert_bind_args idc8
- Raw.ident.Literal;
- args0 <- invert_bind_args idc7
- Raw.ident.Literal;
- args1 <- invert_bind_args idc6
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc5 Raw.ident.Z_land;
- args3 <- invert_bind_args idc4
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc3
- Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc2
- Raw.ident.Z_cast;
- args6 <- invert_bind_args idc1
- Raw.ident.Literal;
- args7 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_get_carry;
- match
- pattern.type.unify_extracted
- ((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args7) -> (projT1 args6)) ->
- (s8 -> (projT1 args0)) ->
- (projT1 args))%ptype
- with
- | Some (_, _, (_, _, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- ((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args7) -> (projT1 args6)) ->
- (s8 -> (projT1 args0)) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args7);
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args6);
- v <- type.try_make_transport_cps s8 ℤ;
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args0);
- xv2 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x10 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s9 xx : Z)
- (rshiftl rland
- ry : zrange)
- (y : expr ℤ)
- (mask offset : Z) =>
- if
- (s9 =? 2 ^ Z.log2 s9) &&
- (ZRange.normalize
- rland <<
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftl)%zrange &&
- (ZRange.normalize ry &'
- ZRange.normalize
- (ZRange.constant
- mask) <=?
- ZRange.normalize
- rland)%zrange &&
- (ZRange.normalize
- rshiftl <=?
- r[0 ~> s9 - 1])%zrange &&
- (mask =?
- Z.ones
- (Z.log2 s9 -
- offset)) &&
- (0 <=? offset) &&
- (offset <=? Z.log2 s9)
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_add
- (Z.log2 s9)
- offset)%expr @
- ((##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x10, _) := xv in
- x10)
- (let (x10, _) :=
- xv0 in
- x10) args5 args3
- args1
- (v
- (Compile.reflect x9))
- (let (x10, _) :=
- xv1 in
- x10)
- (let (x10, _) :=
- xv2 in
- x10);
- Some (Base x10));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ @expr.Abs _
- _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ (_ @ _)))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ @expr.LetIn
- _ _ _ _ _ _ _))%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ ($_)%expr _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (_ @ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ #(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _
- _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _
- _ _ _ @ _))%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _
- _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _
- _ _ _)%expr_pat => None
- | _ => None
- end;;
- match x5 with
- | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t4 idc4)
- x6 =>
- match x4 with
- | @expr.Ident _ _ _ t5 idc5 =>
- args <- invert_bind_args idc5
- Raw.ident.Literal;
- args0 <- invert_bind_args idc4
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc3
- Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc2
- Raw.ident.Z_cast;
- args3 <- invert_bind_args idc1
- Raw.ident.Literal;
- args4 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_get_carry;
- match
- pattern.type.unify_extracted
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> (projT1 args3)) ->
- s5 -> (projT1 args))%ptype
- with
- | Some (_, _, (_, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> (projT1 args3)) ->
- s5 -> (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args4);
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args3);
- v <- type.try_make_transport_cps s5 ℤ;
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x7 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s6 xx : Z)
- (rshiftr ry : zrange)
- (y : expr ℤ)
- (offset : Z) =>
- if
- (s6 =? 2 ^ Z.log2 s6) &&
- (ZRange.normalize ry >>
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftr)%zrange &&
- (ZRange.normalize
- rshiftr <=?
- r[0 ~> s6 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_add
- (Z.log2 s6)
- (- offset))%expr @
- ((##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x7, _) := xv in
- x7)
- (let (x7, _) := xv0 in
- x7) args2 args0
- (v (Compile.reflect x6))
- (let (x7, _) := xv1 in
- x7);
- Some (Base x7));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _
- _ s5 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
- _ s5 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s5 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | _ => None
- end;;
- args <- invert_bind_args idc2 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc1 Raw.ident.Literal;
- args1 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_add_get_carry;
- match
- pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> (projT1 args0)) -> s2)%ptype
- with
- | Some (_, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> (projT1 args0)) -> s2)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s2 ℤ;
- fv <- (x4 <- (let
- '(r1, r2)%zrange := range in
- fun (s3 xx : Z) (ry : zrange)
- (y : expr ℤ) =>
- if
- (s3 =? 2 ^ Z.log2 s3) &&
- (ZRange.normalize ry <=?
- r[0 ~> s3 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_add (Z.log2 s3) 0)%expr @
- ((##xx)%expr,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x4, _) := xv in x4)
- (let (x4, _) := xv0 in x4) args
- (v (Compile.reflect x3));
- Some (Base x4));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s2 _ ($_)%expr _ | @expr.App _ _ _ s2 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s2 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s2 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t1 idc1) x3 =>
+ | @expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t2 idc2) x4 =>
match x0 with
- | @expr.Ident _ _ _ t2 idc2 =>
- match x3 with
- | (@expr.Ident _ _ _ t3 idc3 @ x5 @ x4)%expr_pat =>
- match x5 with
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ @expr.Ident
- _ _ _ t7 idc7))%expr_pat =>
- match x4 with
- | @expr.Ident _ _ _ t8 idc8 =>
- args <- invert_bind_args idc8
- Raw.ident.Literal;
- args0 <- invert_bind_args idc7
- Raw.ident.Literal;
- args1 <- invert_bind_args idc6
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc5 Raw.ident.Z_land;
- args3 <- invert_bind_args idc4
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc3
- Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc2
- Raw.ident.Literal;
- args6 <- invert_bind_args idc1
- Raw.ident.Z_cast;
- args7 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_get_carry;
- match
- pattern.type.unify_extracted
- ((ℤ -> (ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((projT1 args7) ->
- (s8 -> (projT1 args0)) ->
- (projT1 args)) -> (projT1 args5))%ptype
- with
- | Some (_, (_, _, _), _)%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- ((ℤ -> (ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((projT1 args7) ->
- (s8 -> (projT1 args0)) ->
- (projT1 args)) -> (projT1 args5))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args7);
- v <- type.try_make_transport_cps s8 ℤ;
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args0);
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- xv2 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args5);
- fv <- (x10 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s9 : Z)
- (rshiftl rland
- ry : zrange)
- (y : expr ℤ)
- (mask offset xx : Z)
- =>
- if
- (s9 =? 2 ^ Z.log2 s9) &&
- (ZRange.normalize
- rland <<
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftl)%zrange &&
- (ZRange.normalize ry &'
- ZRange.normalize
- (ZRange.constant
- mask) <=?
- ZRange.normalize
- rland)%zrange &&
- (ZRange.normalize
- rshiftl <=?
- r[0 ~> s9 - 1])%zrange &&
- (mask =?
- Z.ones
- (Z.log2 s9 -
- offset)) &&
- (0 <=? offset) &&
- (offset <=? Z.log2 s9)
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_add
- (Z.log2 s9)
- offset)%expr @
- ((##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x10, _) := xv in
- x10) args6 args3
- args1
- (v
- (Compile.reflect x9))
- (let (x10, _) :=
- xv0 in
- x10)
- (let (x10, _) :=
- xv1 in
- x10)
- (let (x10, _) :=
- xv2 in
- x10);
- Some (Base x10));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ @expr.Abs _
- _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ (_ @ _)))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ @expr.LetIn
- _ _ _ _ _ _ _))%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ ($_)%expr _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (_ @ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ #(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _
- _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _
- _ _ _ @ _))%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _
- _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _
- _ _ _)%expr_pat => None
- | _ => None
- end;;
- match x5 with
- | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t4 idc4)
- x6 =>
- match x4 with
- | @expr.Ident _ _ _ t5 idc5 =>
- args <- invert_bind_args idc5
- Raw.ident.Literal;
- args0 <- invert_bind_args idc4
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc3
- Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc2
- Raw.ident.Literal;
- args3 <- invert_bind_args idc1
- Raw.ident.Z_cast;
- args4 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_get_carry;
- match
- pattern.type.unify_extracted
- ((ℤ -> ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args4) -> s5 -> (projT1 args)) ->
- (projT1 args2))%ptype
- with
- | Some (_, (_, _), _)%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- ((ℤ -> ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args4) ->
- s5 -> (projT1 args)) ->
- (projT1 args2))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args4);
- v <- type.try_make_transport_cps s5 ℤ;
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args2);
- fv <- (x7 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s6 : Z)
- (rshiftr ry : zrange)
- (y : expr ℤ)
- (offset xx : Z) =>
- if
- (s6 =? 2 ^ Z.log2 s6) &&
- (ZRange.normalize ry >>
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftr)%zrange &&
- (ZRange.normalize
- rshiftr <=?
- r[0 ~> s6 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_add
- (Z.log2 s6)
- (- offset))%expr @
- ((##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x7, _) := xv in
- x7) args3 args0
- (v (Compile.reflect x6))
- (let (x7, _) := xv0 in
- x7)
- (let (x7, _) := xv1 in
- x7);
- Some (Base x7));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _
- _ s5 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
- _ s5 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s5 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | _ => None
- end;;
- args <- invert_bind_args idc2 Raw.ident.Literal;
- args0 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_add_get_carry;
- match
- pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s2) -> (projT1 args))%ptype
- with
- | Some (_, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s2) -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s2 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x4 <- (let
- '(r1, r2)%zrange := range in
- fun (s3 : Z) (rx : zrange)
- (x4 : expr ℤ) (yy : Z) =>
- if
- (s3 =? 2 ^ Z.log2 s3) &&
- (ZRange.normalize
- (ZRange.constant yy) <=?
- r[0 ~> s3 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_add (Z.log2 s3) 0)%expr @
- (#(Z_cast rx)%expr @ x4,
- (##yy)%expr)))%expr_pat
- else None)
- (let (x4, _) := xv in x4) args0
- (v (Compile.reflect x3))
- (let (x4, _) := xv0 in x4);
- Some (Base x4));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t2 idc2) x4 =>
- match x4 with
- | (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Ident _ _
- _ t7 idc7)) @ @expr.Ident _ _ _ t8 idc8)%expr_pat =>
- args <- invert_bind_args idc8 Raw.ident.Literal;
- args0 <- invert_bind_args idc7 Raw.ident.Literal;
- args1 <- invert_bind_args idc6 Raw.ident.Z_cast;
- _ <- invert_bind_args idc5 Raw.ident.Z_land;
- args3 <- invert_bind_args idc4 Raw.ident.Z_cast;
- _ <- invert_bind_args idc3 Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args6 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args7 <- invert_bind_args idc0 Raw.ident.Literal;
+ | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t3 idc3) x5 =>
+ match x5 with
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @
+ (@expr.Ident _ _ _ t10 idc10 @ @expr.Ident _ _ _ t11
+ idc11))%expr_pat =>
+ args <- invert_bind_args idc11 Raw.ident.Literal;
+ args0 <- invert_bind_args idc10 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc9 Raw.ident.Literal;
+ args2 <- invert_bind_args idc8 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc7 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc6 Raw.ident.Z_land;
+ args5 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc4 Raw.ident.Z_shiftl;
+ args7 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args8 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args9 <- invert_bind_args idc1 Raw.ident.Literal;
+ args10 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_add_get_carry;
match
pattern.type.unify_extracted
((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args7) -> s2) ->
- (s9 -> (projT1 args0)) -> (projT1 args))%ptype
+ (((projT1 args9) -> s3) ->
+ (s10 -> (projT1 args1)) -> (projT1 args))%ptype
with
| Some (_, _, (_, _, _))%zrange =>
if
type.type_beq base.type base.type.type_beq
((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args7) -> s2) ->
- (s9 -> (projT1 args0)) -> (projT1 args))%ptype
+ (((projT1 args9) -> s3) ->
+ (s10 -> (projT1 args1)) -> (projT1 args))%ptype
then
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args7);
- v <- type.try_make_transport_cps s2 ℤ;
- v0 <- type.try_make_transport_cps s9 ℤ;
+ ##(projT2 args9);
+ v <- type.try_make_transport_cps s3 ℤ;
+ v0 <- type.try_make_transport_cps s10 ℤ;
xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
+ ##(projT2 args1);
xv1 <- ident.unify pattern.ident.Literal
##(projT2 args);
- fv <- (x11 <- (let
+ fv <- (x14 <- (let
'(r1, r2)%zrange := range in
- fun (s10 : Z) (rx : zrange)
- (x11 : expr ℤ)
+ fun (rs : zrange) (s13 : Z)
+ (rx : zrange) (x14 : expr ℤ)
(rshiftl rland ry : zrange)
(y : expr ℤ)
- (mask offset : Z) =>
+ (rmask : zrange) (mask : Z)
+ (roffset : zrange)
+ (offset : Z) =>
if
- (s10 =? 2 ^ Z.log2 s10) &&
+ (s13 =? 2 ^ Z.log2 s13) &&
(ZRange.normalize rland <<
ZRange.normalize
(ZRange.constant offset) <=?
@@ -3780,195 +3337,308 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ZRange.constant mask) <=?
ZRange.normalize rland)%zrange &&
(ZRange.normalize rshiftl <=?
- r[0 ~> s10 - 1])%zrange &&
+ r[0 ~> s13 - 1])%zrange &&
(mask =?
- Z.ones (Z.log2 s10 - offset)) &&
+ Z.ones (Z.log2 s13 - offset)) &&
(0 <=? offset) &&
- (offset <=? Z.log2 s10)
+ (offset <=? Z.log2 s13) &&
+ is_bounded_by_bool s13
+ (ZRange.normalize rs) &&
+ is_bounded_by_bool mask
+ (ZRange.normalize rmask) &&
+ is_bounded_by_bool offset
+ (ZRange.normalize roffset)
then
Some
(#(Z_cast2 (r1, r2))%expr @
- (#(fancy_add (Z.log2 s10)
+ (#(fancy_add (Z.log2 s13)
offset)%expr @
- (#(Z_cast rx)%expr @ x11,
+ (#(Z_cast rx)%expr @ x14,
#(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x11, _) := xv in x11)
- args6 (v (Compile.reflect x3))
- args5 args3 args1
- (v0 (Compile.reflect x10))
- (let (x11, _) := xv0 in x11)
- (let (x11, _) := xv1 in x11);
- Some (Base x11));
+ else None) args10
+ (let (x14, _) := xv in x14)
+ args8 (v (Compile.reflect x4))
+ args7 args5 args3
+ (v0 (Compile.reflect x11))
+ args2
+ (let (x14, _) := xv0 in x14)
+ args0
+ (let (x14, _) := xv1 in x14);
+ Some (Base x14));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Ident _ _
- _ t7 idc7)) @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Ident _ _
- _ t7 idc7)) @ @expr.Abs _ _ _ _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Ident _ _
- _ t7 idc7)) @ (_ @ _))%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Ident _ _
- _ t7 idc7)) @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @
+ (@expr.Ident _ _ _ t10 idc10 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @
+ (@expr.Ident _ _ _ t10 idc10 @ @expr.Abs _ _ _ _ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ (@expr.Ident _ _ _ t10 idc10 @ (_ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @
+ (@expr.Ident _ _ _ t10 idc10 @ @expr.LetIn _ _ _ _ _
+ _ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ (_ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
None
- | (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ ($_)%expr)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Abs _ _ _
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ ($_)%expr))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Abs _ _ _ _ _ _))) @
+ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ (_ @ _)))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.LetIn _ _ _ _ _
+ _ _))) @ _)%expr_pat => None
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @ #(_))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @ ($_)%expr)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @ @expr.Abs _ _ _
_ _ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ (_ @ _))) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.LetIn _ _
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @ (($_)%expr @ _))) @
+ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Abs _ _ _ _ _ _ @ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @ (_ @ _ @ _))) @
+ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @ @expr.LetIn _ _
_ _ _ _ _)) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
($_)%expr _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
(@expr.Abs _ _ _ _ _ _) _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
(_ @ _) _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
(@expr.LetIn _ _ _ _ _ _ _) _ @ _)) @ _)%expr_pat =>
None
- | (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ #(_) @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _ @
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _ _ _ @
_)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _ _ _
_ _ @ _)) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ #(_)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _ _) @
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ #(_)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _) @
_)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ (#(_) @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.Abs _ _ _ _ _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _ @ _)) @
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _)) @
_)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.Abs _ _ _ _ _ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.LetIn _ _ _ _ _ _ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.LetIn _ _ _ _ _ _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _ _ _
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _
_) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t3 idc3 @ #(_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ ($_)%expr @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.Abs _ _ _ _ _ _ @
+ | (@expr.Ident _ _ _ t4 idc4 @ #(_) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _ _ @
_)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ (($_)%expr @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
+ (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
(@expr.Abs _ _ _ _ _ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ (_ @ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
+ (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
(@expr.LetIn _ _ _ _ _ _ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.LetIn _ _ _ _ _ _
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _ _ _
_ @ _)%expr_pat => None
| _ => None
end;;
- match x3 with
- | (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Ident _ _
- _ t7 idc7)) @ @expr.Ident _ _ _ t8 idc8)%expr_pat =>
- args <- invert_bind_args idc8 Raw.ident.Literal;
- args0 <- invert_bind_args idc7 Raw.ident.Literal;
- args1 <- invert_bind_args idc6 Raw.ident.Z_cast;
- _ <- invert_bind_args idc5 Raw.ident.Z_land;
- args3 <- invert_bind_args idc4 Raw.ident.Z_cast;
- _ <- invert_bind_args idc3 Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args6 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args7 <- invert_bind_args idc0 Raw.ident.Literal;
+ match x4 with
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @
+ (@expr.Ident _ _ _ t10 idc10 @ @expr.Ident _ _ _ t11
+ idc11))%expr_pat =>
+ args <- invert_bind_args idc11 Raw.ident.Literal;
+ args0 <- invert_bind_args idc10 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc9 Raw.ident.Literal;
+ args2 <- invert_bind_args idc8 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc7 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc6 Raw.ident.Z_land;
+ args5 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc4 Raw.ident.Z_shiftl;
+ args7 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args8 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args9 <- invert_bind_args idc1 Raw.ident.Literal;
+ args10 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_add_get_carry;
match
pattern.type.unify_extracted
((ℤ -> (ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((projT1 args7) ->
- (s9 -> (projT1 args0)) -> (projT1 args)) ->
- s3)%ptype
+ (((projT1 args9) ->
+ (s10 -> (projT1 args1)) -> (projT1 args)) ->
+ s4)%ptype
with
| Some (_, (_, _, _), _)%zrange =>
if
type.type_beq base.type base.type.type_beq
((ℤ -> (ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- (((projT1 args7) ->
- (s9 -> (projT1 args0)) -> (projT1 args)) ->
- s3)%ptype
+ (((projT1 args9) ->
+ (s10 -> (projT1 args1)) -> (projT1 args)) ->
+ s4)%ptype
then
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args7);
- v <- type.try_make_transport_cps s9 ℤ;
+ ##(projT2 args9);
+ v <- type.try_make_transport_cps s10 ℤ;
xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
+ ##(projT2 args1);
xv1 <- ident.unify pattern.ident.Literal
##(projT2 args);
- v0 <- type.try_make_transport_cps s3 ℤ;
- fv <- (x11 <- (let
+ v0 <- type.try_make_transport_cps s4 ℤ;
+ fv <- (x14 <- (let
'(r1, r2)%zrange := range in
- fun (s10 : Z)
+ fun (rs : zrange) (s13 : Z)
(rshiftl rland ry : zrange)
(y : expr ℤ)
- (mask offset : Z)
- (rx : zrange) (x11 : expr ℤ)
- =>
+ (rmask : zrange) (mask : Z)
+ (roffset : zrange)
+ (offset : Z) (rx : zrange)
+ (x14 : expr ℤ) =>
if
- (s10 =? 2 ^ Z.log2 s10) &&
+ (s13 =? 2 ^ Z.log2 s13) &&
(ZRange.normalize rland <<
ZRange.normalize
(ZRange.constant offset) <=?
@@ -3978,418 +3648,607 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ZRange.constant mask) <=?
ZRange.normalize rland)%zrange &&
(ZRange.normalize rshiftl <=?
- r[0 ~> s10 - 1])%zrange &&
+ r[0 ~> s13 - 1])%zrange &&
(mask =?
- Z.ones (Z.log2 s10 - offset)) &&
+ Z.ones (Z.log2 s13 - offset)) &&
(0 <=? offset) &&
- (offset <=? Z.log2 s10)
+ (offset <=? Z.log2 s13) &&
+ is_bounded_by_bool s13
+ (ZRange.normalize rs) &&
+ is_bounded_by_bool mask
+ (ZRange.normalize rmask) &&
+ is_bounded_by_bool offset
+ (ZRange.normalize roffset)
then
Some
(#(Z_cast2 (r1, r2))%expr @
- (#(fancy_add (Z.log2 s10)
+ (#(fancy_add (Z.log2 s13)
offset)%expr @
- (#(Z_cast rx)%expr @ x11,
+ (#(Z_cast rx)%expr @ x14,
#(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x11, _) := xv in x11)
- args6 args3 args1
- (v (Compile.reflect x10))
- (let (x11, _) := xv0 in x11)
- (let (x11, _) := xv1 in x11)
- args5
- (v0 (Compile.reflect x4));
- Some (Base x11));
+ else None) args10
+ (let (x14, _) := xv in x14)
+ args8 args5 args3
+ (v (Compile.reflect x11))
+ args2
+ (let (x14, _) := xv0 in x14)
+ args0
+ (let (x14, _) := xv1 in x14)
+ args7
+ (v0 (Compile.reflect x5));
+ Some (Base x14));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Ident _ _
- _ t7 idc7)) @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Ident _ _
- _ t7 idc7)) @ @expr.Abs _ _ _ _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Ident _ _
- _ t7 idc7)) @ (_ @ _))%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Ident _ _
- _ t7 idc7)) @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @
+ (@expr.Ident _ _ _ t10 idc10 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @
+ (@expr.Ident _ _ _ t10 idc10 @ @expr.Abs _ _ _ _ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ (@expr.Ident _ _ _ t10 idc10 @ (_ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @
+ (@expr.Ident _ _ _ t10 idc10 @ @expr.LetIn _ _ _ _ _
+ _ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ (_ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _ t9
+ idc9))) @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
None
- | (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ ($_)%expr)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Abs _ _ _
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ ($_)%expr))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Abs _ _ _ _ _ _))) @
+ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ (_ @ _)))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.LetIn _ _ _ _ _
+ _ _))) @ _)%expr_pat => None
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @ #(_))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @ ($_)%expr)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @ @expr.Abs _ _ _
_ _ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ (_ @ _))) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
- (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.LetIn _ _
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @ (($_)%expr @ _))) @
+ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Abs _ _ _ _ _ _ @ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @ (_ @ _ @ _))) @
+ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
+ (@expr.Ident _ _ _ t7 idc7) x11 @ @expr.LetIn _ _
_ _ _ _ _)) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
($_)%expr _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
(@expr.Abs _ _ _ _ _ _) _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
(_ @ _) _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10 _
(@expr.LetIn _ _ _ _ _ _ _) _ @ _)) @ _)%expr_pat =>
None
- | (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ #(_) @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _ @
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _ _ _ @
_)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _ _ _
_ _ @ _)) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ #(_)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _ _) @
+ | (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ #(_)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _ _) @
_)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ (#(_) @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.Abs _ _ _ _ _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _ @ _)) @
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _)) @
_)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.Abs _ _ _ _ _ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.LetIn _ _ _ _ _ _ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.LetIn _ _ _ _ _ _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _ _ _
+ (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _ _ _
_) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t3 idc3 @ #(_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ ($_)%expr @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.Abs _ _ _ _ _ _ @
+ | (@expr.Ident _ _ _ t4 idc4 @ #(_) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _ _ @
_)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ (($_)%expr @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
+ (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
(@expr.Abs _ _ _ _ _ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ (_ @ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @
+ (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @
(@expr.LetIn _ _ _ _ _ _ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.LetIn _ _ _ _ _ _
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _ _ _
_ @ _)%expr_pat => None
| _ => None
end;;
- match x4 with
- | (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
- (@expr.Ident _ _ _ t4 idc4) x7 @ @expr.Ident _ _ _
- t5 idc5)%expr_pat =>
- args <- invert_bind_args idc5 Raw.ident.Literal;
- args0 <- invert_bind_args idc4 Raw.ident.Z_cast;
- _ <- invert_bind_args idc3 Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args4 <- invert_bind_args idc0 Raw.ident.Literal;
+ match x5 with
+ | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Ident _ _ _ t7
+ idc7))%expr_pat =>
+ args <- invert_bind_args idc7 Raw.ident.Literal;
+ args0 <- invert_bind_args idc6 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc4 Raw.ident.Z_shiftr;
+ args3 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args5 <- invert_bind_args idc1 Raw.ident.Literal;
+ args6 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_add_get_carry;
match
pattern.type.unify_extracted
((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> s2) -> s6 -> (projT1 args))%ptype
+ (((projT1 args5) -> s3) -> s7 -> (projT1 args))%ptype
with
| Some (_, _, (_, _))%zrange =>
if
type.type_beq base.type base.type.type_beq
((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> s2) ->
- s6 -> (projT1 args))%ptype
+ (((projT1 args5) -> s3) ->
+ s7 -> (projT1 args))%ptype
then
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args4);
- v <- type.try_make_transport_cps s2 ℤ;
- v0 <- type.try_make_transport_cps s6 ℤ;
+ ##(projT2 args5);
+ v <- type.try_make_transport_cps s3 ℤ;
+ v0 <- type.try_make_transport_cps s7 ℤ;
xv0 <- ident.unify pattern.ident.Literal
##(projT2 args);
- fv <- (x8 <- (let
- '(r1, r2)%zrange := range in
- fun (s7 : Z) (rx : zrange)
- (x8 : expr ℤ)
- (rshiftr ry : zrange)
- (y : expr ℤ) (offset : Z) =>
- if
- (s7 =? 2 ^ Z.log2 s7) &&
- (ZRange.normalize ry >>
- ZRange.normalize
- (ZRange.constant offset) <=?
- ZRange.normalize rshiftr)%zrange &&
- (ZRange.normalize rshiftr <=?
- r[0 ~> s7 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_add (Z.log2 s7)
- (- offset))%expr @
- (#(Z_cast rx)%expr @ x8,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x8, _) := xv in x8) args3
- (v (Compile.reflect x3)) args2
- args0 (v0 (Compile.reflect x7))
- (let (x8, _) := xv0 in x8);
- Some (Base x8));
+ fv <- (x10 <- (let
+ '(r1, r2)%zrange := range in
+ fun (rs : zrange) (s9 : Z)
+ (rx : zrange) (x10 : expr ℤ)
+ (rshiftr ry : zrange)
+ (y : expr ℤ)
+ (roffset : zrange)
+ (offset : Z) =>
+ if
+ (s9 =? 2 ^ Z.log2 s9) &&
+ (ZRange.normalize ry >>
+ ZRange.normalize
+ (ZRange.constant offset) <=?
+ ZRange.normalize rshiftr)%zrange &&
+ (ZRange.normalize rshiftr <=?
+ r[0 ~> s9 - 1])%zrange &&
+ is_bounded_by_bool s9
+ (ZRange.normalize rs) &&
+ is_bounded_by_bool offset
+ (ZRange.normalize roffset)
+ then
+ Some
+ (#(Z_cast2 (r1, r2))%expr @
+ (#(fancy_add (Z.log2 s9)
+ (- offset))%expr @
+ (#(Z_cast rx)%expr @ x10,
+ #(Z_cast ry)%expr @ y)))%expr_pat
+ else None) args6
+ (let (x10, _) := xv in x10)
+ args4 (v (Compile.reflect x4))
+ args3 args1
+ (v0 (Compile.reflect x8))
+ args0
+ (let (x10, _) := xv0 in x10);
+ Some (Base x10));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
- (@expr.Ident _ _ _ t4 idc4) x7 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
- (@expr.Ident _ _ _ t4 idc4) x7 @ @expr.Abs _ _ _ _ _
+ | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.Ident _ _ _ t6 idc6 @ (_ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _ _ _ _
+ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @ @expr.Abs _ _ _ _ _
_)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
- (@expr.Ident _ _ _ t4 idc4) x7 @ (_ @ _))%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
- (@expr.Ident _ _ _ t4 idc4) x7 @ @expr.LetIn _ _ _ _
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @ (_ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @ @expr.LetIn _ _ _ _
_ _ _)%expr_pat => None
- | (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
($_)%expr _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
(@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
(_ @ _) _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
(@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t3 idc3 @ #(_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ ($_)%expr @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.Abs _ _ _ _ _ _ @
+ | (@expr.Ident _ _ _ t4 idc4 @ #(_) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _ _ @
_)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.LetIn _ _ _ _ _ _
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _ _ _
_ @ _)%expr_pat => None
| _ => None
end;;
- match x3 with
- | (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
- (@expr.Ident _ _ _ t4 idc4) x7 @ @expr.Ident _ _ _
- t5 idc5)%expr_pat =>
- args <- invert_bind_args idc5 Raw.ident.Literal;
- args0 <- invert_bind_args idc4 Raw.ident.Z_cast;
- _ <- invert_bind_args idc3 Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args4 <- invert_bind_args idc0 Raw.ident.Literal;
+ match x4 with
+ | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Ident _ _ _ t7
+ idc7))%expr_pat =>
+ args <- invert_bind_args idc7 Raw.ident.Literal;
+ args0 <- invert_bind_args idc6 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc5 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc4 Raw.ident.Z_shiftr;
+ args3 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args5 <- invert_bind_args idc1 Raw.ident.Literal;
+ args6 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_add_get_carry;
match
pattern.type.unify_extracted
((ℤ -> ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args4) -> s6 -> (projT1 args)) -> s3)%ptype
+ (((projT1 args5) -> s7 -> (projT1 args)) -> s4)%ptype
with
| Some (_, (_, _), _)%zrange =>
if
type.type_beq base.type base.type.type_beq
((ℤ -> ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args4) -> s6 -> (projT1 args)) ->
- s3)%ptype
+ (((projT1 args5) -> s7 -> (projT1 args)) ->
+ s4)%ptype
then
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args4);
- v <- type.try_make_transport_cps s6 ℤ;
+ ##(projT2 args5);
+ v <- type.try_make_transport_cps s7 ℤ;
xv0 <- ident.unify pattern.ident.Literal
##(projT2 args);
- v0 <- type.try_make_transport_cps s3 ℤ;
- fv <- (x8 <- (let
- '(r1, r2)%zrange := range in
- fun (s7 : Z)
- (rshiftr ry : zrange)
- (y : expr ℤ) (offset : Z)
- (rx : zrange) (x8 : expr ℤ)
- =>
- if
- (s7 =? 2 ^ Z.log2 s7) &&
- (ZRange.normalize ry >>
- ZRange.normalize
- (ZRange.constant offset) <=?
- ZRange.normalize rshiftr)%zrange &&
- (ZRange.normalize rshiftr <=?
- r[0 ~> s7 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_add (Z.log2 s7)
- (- offset))%expr @
- (#(Z_cast rx)%expr @ x8,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x8, _) := xv in x8) args3
- args0 (v (Compile.reflect x7))
- (let (x8, _) := xv0 in x8)
- args2 (v0 (Compile.reflect x4));
- Some (Base x8));
+ v0 <- type.try_make_transport_cps s4 ℤ;
+ fv <- (x10 <- (let
+ '(r1, r2)%zrange := range in
+ fun (rs : zrange) (s9 : Z)
+ (rshiftr ry : zrange)
+ (y : expr ℤ)
+ (roffset : zrange)
+ (offset : Z) (rx : zrange)
+ (x10 : expr ℤ) =>
+ if
+ (s9 =? 2 ^ Z.log2 s9) &&
+ (ZRange.normalize ry >>
+ ZRange.normalize
+ (ZRange.constant offset) <=?
+ ZRange.normalize rshiftr)%zrange &&
+ (ZRange.normalize rshiftr <=?
+ r[0 ~> s9 - 1])%zrange &&
+ is_bounded_by_bool s9
+ (ZRange.normalize rs) &&
+ is_bounded_by_bool offset
+ (ZRange.normalize roffset)
+ then
+ Some
+ (#(Z_cast2 (r1, r2))%expr @
+ (#(fancy_add (Z.log2 s9)
+ (- offset))%expr @
+ (#(Z_cast rx)%expr @ x10,
+ #(Z_cast ry)%expr @ y)))%expr_pat
+ else None) args6
+ (let (x10, _) := xv in x10)
+ args4 args1
+ (v (Compile.reflect x8)) args0
+ (let (x10, _) := xv0 in x10)
+ args3
+ (v0 (Compile.reflect x5));
+ Some (Base x10));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
- (@expr.Ident _ _ _ t4 idc4) x7 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
- (@expr.Ident _ _ _ t4 idc4) x7 @ @expr.Abs _ _ _ _ _
+ | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.Ident _ _ _ t6 idc6 @ (_ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _ _ _ _
+ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @ @expr.Abs _ _ _ _ _
_)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
- (@expr.Ident _ _ _ t4 idc4) x7 @ (_ @ _))%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
- (@expr.Ident _ _ _ t4 idc4) x7 @ @expr.LetIn _ _ _ _
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @ (_ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
+ (@expr.Ident _ _ _ t5 idc5) x8 @ @expr.LetIn _ _ _ _
_ _ _)%expr_pat => None
- | (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
($_)%expr _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
(@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
(_ @ _) _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.App _ _ _ s6 _
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
(@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t3 idc3 @ #(_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ ($_)%expr @ _)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.Abs _ _ _ _ _ _ @
+ | (@expr.Ident _ _ _ t4 idc4 @ #(_) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr @ _)%expr_pat |
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _ _ @
_)%expr_pat |
- (@expr.Ident _ _ _ t3 idc3 @ @expr.LetIn _ _ _ _ _ _
+ (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _ _ _
_ @ _)%expr_pat => None
| _ => None
end;;
- args <- invert_bind_args idc2 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc0 Raw.ident.Literal;
+ args <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc1 Raw.ident.Literal;
+ args2 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_add_get_carry;
match
pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s2) -> s3)%ptype
+ (((projT1 args1) -> s3) -> s4)%ptype
with
| Some (_, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s2) -> s3)%ptype
+ (((projT1 args1) -> s3) -> s4)%ptype
then
xv <- ident.unify pattern.ident.Literal
##(projT2 args1);
- v <- type.try_make_transport_cps s2 ℤ;
- v0 <- type.try_make_transport_cps s3 ℤ;
- fv <- (x5 <- (let
+ v <- type.try_make_transport_cps s3 ℤ;
+ v0 <- type.try_make_transport_cps s4 ℤ;
+ fv <- (x6 <- (let
'(r1, r2)%zrange := range in
- fun (s4 : Z) (rx : zrange)
- (x5 : expr ℤ) (ry : zrange)
- (y : expr ℤ) =>
+ fun (rs : zrange) (s5 : Z)
+ (rx : zrange) (x6 : expr ℤ)
+ (ry : zrange) (y : expr ℤ) =>
if
- (s4 =? 2 ^ Z.log2 s4) &&
+ (s5 =? 2 ^ Z.log2 s5) &&
(ZRange.normalize ry <=?
- r[0 ~> s4 - 1])%zrange
+ r[0 ~> s5 - 1])%zrange &&
+ is_bounded_by_bool s5
+ (ZRange.normalize rs)
then
Some
(#(Z_cast2 (r1, r2))%expr @
- (#(fancy_add (Z.log2 s4) 0)%expr @
- (#(Z_cast rx)%expr @ x5,
+ (#(fancy_add (Z.log2 s5) 0)%expr @
+ (#(Z_cast rx)%expr @ x6,
#(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x5, _) := xv in x5) args0
- (v (Compile.reflect x3)) args
- (v0 (Compile.reflect x4));
- Some (Base x5));
+ else None) args2
+ (let (x6, _) := xv in x6) args0
+ (v (Compile.reflect x4)) args
+ (v0 (Compile.reflect x5));
+ Some (Base x6));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | @expr.App _ _ _ s3 _ ($_)%expr _ | @expr.App _ _ _ s3 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s3 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s3 _
+ | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s4 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
- | @expr.App _ _ _ s2 _ ($_)%expr _ | @expr.App _ _ _ s2 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s2 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s2 _
+ | @expr.App _ _ _ s3 _ ($_)%expr _ | @expr.App _ _ _ s3 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s3 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s3 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
+ | (@expr.Ident _ _ _ t0 idc0 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end;;
match x2 with
- | @expr.Ident _ _ _ t0 idc0 =>
+ | (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat =>
match x1 with
- | @expr.Ident _ _ _ t1 idc1 =>
+ | @expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t2 idc2) x4 =>
match x0 with
- | @expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t2 idc2) x3 =>
- match x3 with
- | (@expr.Ident _ _ _ t3 idc3 @ x5 @ x4)%expr_pat =>
- match x5 with
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ @expr.Ident
- _ _ _ t7 idc7))%expr_pat =>
- match x4 with
- | @expr.Ident _ _ _ t8 idc8 =>
- args <- invert_bind_args idc8
+ | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t3 idc3) x5 =>
+ match x5 with
+ | (@expr.Ident _ _ _ t4 idc4 @ x7 @ x6)%expr_pat =>
+ match x7 with
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
+ _ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Ident _ _ _
+ t9 idc9)))%expr_pat =>
+ match x6 with
+ | (@expr.Ident _ _ _ t10 idc10 @ @expr.Ident _
+ _ _ t11 idc11)%expr_pat =>
+ args <- invert_bind_args idc11
Raw.ident.Literal;
- args0 <- invert_bind_args idc7
+ args0 <- invert_bind_args idc10
+ Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc9
Raw.ident.Literal;
- args1 <- invert_bind_args idc6
+ args2 <- invert_bind_args idc8
+ Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc7
Raw.ident.Z_cast;
- _ <- invert_bind_args idc5 Raw.ident.Z_land;
- args3 <- invert_bind_args idc4
+ _ <- invert_bind_args idc6 Raw.ident.Z_land;
+ args5 <- invert_bind_args idc5
Raw.ident.Z_cast;
- _ <- invert_bind_args idc3
+ _ <- invert_bind_args idc4
Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc2
+ args7 <- invert_bind_args idc3
Raw.ident.Z_cast;
- args6 <- invert_bind_args idc1
- Raw.ident.Literal;
- args7 <- invert_bind_args idc0
+ args8 <- invert_bind_args idc2
+ Raw.ident.Z_cast;
+ args9 <- invert_bind_args idc1
Raw.ident.Literal;
+ args10 <- invert_bind_args idc0
+ Raw.ident.Z_cast;
_ <- invert_bind_args idc
Raw.ident.Z_sub_get_borrow;
match
pattern.type.unify_extracted
((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args7) -> (projT1 args6)) ->
- (s8 -> (projT1 args0)) ->
+ (((projT1 args9) -> s3) ->
+ (s10 -> (projT1 args1)) ->
(projT1 args))%ptype
with
| Some (_, _, (_, _, _))%zrange =>
@@ -4397,33 +4256,38 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
type.type_beq base.type
base.type.type_beq
((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args7) -> (projT1 args6)) ->
- (s8 -> (projT1 args0)) ->
+ (((projT1 args9) -> s3) ->
+ (s10 -> (projT1 args1)) ->
(projT1 args))%ptype
then
xv <- ident.unify
pattern.ident.Literal
- ##(projT2 args7);
+ ##(projT2 args9);
+ v <- type.try_make_transport_cps s3 ℤ;
+ v0 <- type.try_make_transport_cps s10 ℤ;
xv0 <- ident.unify
pattern.ident.Literal
- ##(projT2 args6);
- v <- type.try_make_transport_cps s8 ℤ;
+ ##(projT2 args1);
xv1 <- ident.unify
pattern.ident.Literal
- ##(projT2 args0);
- xv2 <- ident.unify
- pattern.ident.Literal
##(projT2 args);
- fv <- (x10 <- (let
+ fv <- (x14 <- (let
'(r1, r2)%zrange :=
range in
- fun (s9 xx : Z)
+ fun (rs : zrange)
+ (s13 : Z)
+ (rx : zrange)
+ (x14 : expr ℤ)
(rshiftl rland
ry : zrange)
(y : expr ℤ)
- (mask offset : Z) =>
+ (rmask : zrange)
+ (mask : Z)
+ (roffset : zrange)
+ (offset : Z) =>
if
- (s9 =? 2 ^ Z.log2 s9) &&
+ (s13 =?
+ 2 ^ Z.log2 s13) &&
(ZRange.normalize
rland <<
ZRange.normalize
@@ -4433,7 +4297,7 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
rshiftl)%zrange &&
(ZRange.normalize
rshiftl <=?
- r[0 ~> s9 - 1])%zrange &&
+ r[0 ~> s13 - 1])%zrange &&
(ZRange.normalize ry &'
ZRange.normalize
(ZRange.constant
@@ -4442,2988 +4306,416 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
rland)%zrange &&
(mask =?
Z.ones
- (Z.log2 s9 -
+ (Z.log2 s13 -
offset)) &&
(0 <=? offset) &&
- (offset <=? Z.log2 s9)
+ (offset <=?
+ Z.log2 s13) &&
+ is_bounded_by_bool
+ s13
+ (ZRange.normalize
+ rs) &&
+ is_bounded_by_bool
+ mask
+ (ZRange.normalize
+ rmask) &&
+ is_bounded_by_bool
+ offset
+ (ZRange.normalize
+ roffset)
then
Some
(#(Z_cast2 (r1, r2))%expr @
(#(fancy_sub
- (Z.log2 s9)
+ (Z.log2 s13)
offset)%expr @
- ((##xx)%expr,
+ (#(Z_cast rx)%expr @
+ x14,
#(Z_cast ry)%expr @
y)))%expr_pat
- else None)
- (let (x10, _) := xv in
- x10)
- (let (x10, _) :=
- xv0 in
- x10) args5 args3
- args1
+ else None) args10
+ (let (x14, _) := xv in
+ x14) args8
(v
- (Compile.reflect x9))
- (let (x10, _) :=
+ (Compile.reflect x4))
+ args7 args5 args3
+ (v0
+ (Compile.reflect
+ x11)) args2
+ (let (x14, _) :=
+ xv0 in
+ x14) args0
+ (let (x14, _) :=
xv1 in
- x10)
- (let (x10, _) :=
- xv2 in
- x10);
- Some (Base x10));
+ x14);
+ Some (Base x14));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
+ | (@expr.Ident _ _ _ t10 idc10 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t10 idc10 @ @expr.Abs _ _
+ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t10 idc10 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t10 idc10 @ @expr.LetIn _
+ _ _ _ _ _ _)%expr_pat => None
| _ => None
end
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ @expr.Abs _
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
+ _ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ ($_)%expr)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
+ _ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.Abs _ _ _ _
+ _ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
+ _ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ (_ @ _))))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
+ _ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Ident _ _ _ t8 idc8 @ @expr.LetIn _ _ _
+ _ _ _ _)))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
+ _ (@expr.Ident _ _ _ t7 idc7) x11 @ #(_)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
+ _ (@expr.Ident _ _ _ t7 idc7) x11 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
+ _ (@expr.Ident _ _ _ t7 idc7) x11 @ @expr.Abs _
_ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ (_ @ _)))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
- _ (@expr.Ident _ _ _ t6 idc6) x9 @ @expr.LetIn
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
+ _ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (($_)%expr @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
+ _ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.Abs _ _ _ _ _ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
+ _ (@expr.Ident _ _ _ t7 idc7) x11 @ (_ @ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
+ _ (@expr.Ident _ _ _ t7 idc7) x11 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
+ _ (@expr.Ident _ _ _ t7 idc7) x11 @ @expr.LetIn
_ _ _ _ _ _ _))%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
_ ($_)%expr _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
_ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
_ (_ @ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s8
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _ s10
_ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ #(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _ _
_ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _ _
_ _ _ @ _))%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _
+ | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _
_)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _
_ _ _)%expr_pat => None
| _ => None
end;;
- match x5 with
- | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t4 idc4)
- x6 =>
- match x4 with
- | @expr.Ident _ _ _ t5 idc5 =>
- args <- invert_bind_args idc5
+ match x7 with
+ | @expr.App _ _ _ s7 _ (@expr.Ident _ _ _ t5 idc5)
+ x8 =>
+ match x6 with
+ | (@expr.Ident _ _ _ t6 idc6 @ @expr.Ident _ _
+ _ t7 idc7)%expr_pat =>
+ args <- invert_bind_args idc7
Raw.ident.Literal;
- args0 <- invert_bind_args idc4
+ args0 <- invert_bind_args idc6
+ Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc5
Raw.ident.Z_cast;
- _ <- invert_bind_args idc3
+ _ <- invert_bind_args idc4
Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc2
+ args3 <- invert_bind_args idc3
Raw.ident.Z_cast;
- args3 <- invert_bind_args idc1
- Raw.ident.Literal;
- args4 <- invert_bind_args idc0
+ args4 <- invert_bind_args idc2
+ Raw.ident.Z_cast;
+ args5 <- invert_bind_args idc1
Raw.ident.Literal;
+ args6 <- invert_bind_args idc0
+ Raw.ident.Z_cast;
_ <- invert_bind_args idc
Raw.ident.Z_sub_get_borrow;
match
pattern.type.unify_extracted
((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> (projT1 args3)) ->
- s5 -> (projT1 args))%ptype
+ (((projT1 args5) -> s3) ->
+ s7 -> (projT1 args))%ptype
with
| Some (_, _, (_, _))%zrange =>
if
type.type_beq base.type
base.type.type_beq
((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> (projT1 args3)) ->
- s5 -> (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args4);
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args3);
- v <- type.try_make_transport_cps s5 ℤ;
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x7 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s6 xx : Z)
- (rshiftr ry : zrange)
- (y : expr ℤ)
- (offset : Z) =>
- if
- (s6 =? 2 ^ Z.log2 s6) &&
- (ZRange.normalize ry >>
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftr)%zrange &&
- (ZRange.normalize
- rshiftr <=?
- r[0 ~> s6 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_sub
- (Z.log2 s6)
- (- offset))%expr @
- ((##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x7, _) := xv in
- x7)
- (let (x7, _) := xv0 in
- x7) args2 args0
- (v (Compile.reflect x6))
- (let (x7, _) := xv1 in
- x7);
- Some (Base x7));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _
- _ s5 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
- _ s5 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s5 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | _ => None
- end;;
- args <- invert_bind_args idc2 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc1 Raw.ident.Literal;
- args1 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_sub_get_borrow;
- match
- pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> (projT1 args0)) -> s2)%ptype
- with
- | Some (_, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> (projT1 args0)) -> s2)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s2 ℤ;
- fv <- (x4 <- (let
- '(r1, r2)%zrange := range in
- fun (s3 xx : Z) (ry : zrange)
- (y : expr ℤ) =>
- if
- (s3 =? 2 ^ Z.log2 s3) &&
- (ZRange.normalize ry <=?
- r[0 ~> s3 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_sub (Z.log2 s3) 0)%expr @
- ((##xx)%expr,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x4, _) := xv in x4)
- (let (x4, _) := xv0 in x4) args
- (v (Compile.reflect x3));
- Some (Base x4));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s2 _ ($_)%expr _ | @expr.App _ _ _ s2 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s2 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s2 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s2 _ (@expr.Ident _ _ _ t1 idc1) x3 =>
- match x0 with
- | @expr.Ident _ _ _ t2 idc2 =>
- args <- invert_bind_args idc2 Raw.ident.Literal;
- args0 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc Raw.ident.Z_sub_get_borrow;
- match
- pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s2) -> (projT1 args))%ptype
- with
- | Some (_, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s2) -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s2 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x4 <- (let
- '(r1, r2)%zrange := range in
- fun (s3 : Z) (rx : zrange)
- (x4 : expr ℤ) (yy : Z) =>
- if
- (s3 =? 2 ^ Z.log2 s3) &&
- (ZRange.normalize
- (ZRange.constant yy) <=?
- r[0 ~> s3 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_sub (Z.log2 s3) 0)%expr @
- (#(Z_cast rx)%expr @ x4,
- (##yy)%expr)))%expr_pat
- else None)
- (let (x4, _) := xv in x4) args0
- (v (Compile.reflect x3))
- (let (x4, _) := xv0 in x4);
- Some (Base x4));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t2 idc2) x4 =>
- match x4 with
- | (@expr.Ident _ _ _ t3 idc3 @ x6 @ x5)%expr_pat =>
- match x6 with
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9
- _ (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Ident
- _ _ _ t7 idc7))%expr_pat =>
- match x5 with
- | @expr.Ident _ _ _ t8 idc8 =>
- args <- invert_bind_args idc8
- Raw.ident.Literal;
- args0 <- invert_bind_args idc7
- Raw.ident.Literal;
- args1 <- invert_bind_args idc6
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc5 Raw.ident.Z_land;
- args3 <- invert_bind_args idc4
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc3
- Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc2
- Raw.ident.Z_cast;
- args6 <- invert_bind_args idc1
- Raw.ident.Z_cast;
- args7 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_get_borrow;
- match
- pattern.type.unify_extracted
- ((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args7) -> s2) ->
- (s9 -> (projT1 args0)) ->
- (projT1 args))%ptype
- with
- | Some (_, _, (_, _, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- ((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args7) -> s2) ->
- (s9 -> (projT1 args0)) ->
- (projT1 args))%ptype
+ (((projT1 args5) -> s3) ->
+ s7 -> (projT1 args))%ptype
then
xv <- ident.unify
pattern.ident.Literal
- ##(projT2 args7);
- v <- type.try_make_transport_cps s2 ℤ;
- v0 <- type.try_make_transport_cps s9 ℤ;
+ ##(projT2 args5);
+ v <- type.try_make_transport_cps s3 ℤ;
+ v0 <- type.try_make_transport_cps s7 ℤ;
xv0 <- ident.unify
pattern.ident.Literal
- ##(projT2 args0);
- xv1 <- ident.unify
- pattern.ident.Literal
##(projT2 args);
- fv <- (x11 <- (let
+ fv <- (x10 <- (let
'(r1, r2)%zrange :=
range in
- fun (s10 : Z)
+ fun (rs : zrange)
+ (s9 : Z)
(rx : zrange)
- (x11 : expr ℤ)
- (rshiftl rland
+ (x10 : expr ℤ)
+ (rshiftr
ry : zrange)
(y : expr ℤ)
- (mask offset : Z) =>
+ (roffset : zrange)
+ (offset : Z) =>
if
- (s10 =?
- 2 ^ Z.log2 s10) &&
- (ZRange.normalize
- rland <<
+ (s9 =? 2 ^ Z.log2 s9) &&
+ (ZRange.normalize ry >>
ZRange.normalize
(ZRange.constant
offset) <=?
ZRange.normalize
- rshiftl)%zrange &&
+ rshiftr)%zrange &&
(ZRange.normalize
- rshiftl <=?
- r[0 ~> s10 - 1])%zrange &&
- (ZRange.normalize ry &'
- ZRange.normalize
- (ZRange.constant
- mask) <=?
- ZRange.normalize
- rland)%zrange &&
- (mask =?
- Z.ones
- (Z.log2 s10 -
- offset)) &&
- (0 <=? offset) &&
- (offset <=?
- Z.log2 s10)
+ rshiftr <=?
+ r[0 ~> s9 - 1])%zrange &&
+ is_bounded_by_bool s9
+ (ZRange.normalize
+ rs) &&
+ is_bounded_by_bool
+ offset
+ (ZRange.normalize
+ roffset)
then
Some
(#(Z_cast2 (r1, r2))%expr @
(#(fancy_sub
- (Z.log2 s10)
- offset)%expr @
+ (Z.log2 s9)
+ (- offset))%expr @
(#(Z_cast rx)%expr @
- x11,
+ x10,
#(Z_cast ry)%expr @
y)))%expr_pat
- else None)
- (let (x11, _) := xv in
- x11) args6
+ else None) args6
+ (let (x10, _) := xv in
+ x10) args4
(v
- (Compile.reflect x3))
- args5 args3 args1
+ (Compile.reflect x4))
+ args3 args1
(v0
- (Compile.reflect
- x10))
- (let (x11, _) :=
+ (Compile.reflect x8))
+ args0
+ (let (x10, _) :=
xv0 in
- x11)
- (let (x11, _) :=
- xv1 in
- x11);
- Some (Base x11));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9
- _ (@expr.Ident _ _ _ t6 idc6) x10 @ ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9
- _ (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Abs _
- _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9
- _ (@expr.Ident _ _ _ t6 idc6) x10 @ (_ @ _)))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9
- _ (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.LetIn
- _ _ _ _ _ _ _))%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9
- _ ($_)%expr _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9
- _ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9
- _ (_ @ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9
- _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ #(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _
- _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _
- _ _ _ @ _))%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _
- _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _
- _ _ _)%expr_pat => None
- | _ => None
- end;;
- match x6 with
- | @expr.App _ _ _ s6 _ (@expr.Ident _ _ _ t4 idc4)
- x7 =>
- match x5 with
- | @expr.Ident _ _ _ t5 idc5 =>
- args <- invert_bind_args idc5
- Raw.ident.Literal;
- args0 <- invert_bind_args idc4
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc3
- Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc2
- Raw.ident.Z_cast;
- args3 <- invert_bind_args idc1
- Raw.ident.Z_cast;
- args4 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_get_borrow;
- match
- pattern.type.unify_extracted
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> s2) ->
- s6 -> (projT1 args))%ptype
- with
- | Some (_, _, (_, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- ((ℤ -> ℤ) -> ℤ -> ℤ)%ptype
- (((projT1 args4) -> s2) ->
- s6 -> (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args4);
- v <- type.try_make_transport_cps s2 ℤ;
- v0 <- type.try_make_transport_cps s6 ℤ;
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x8 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s7 : Z)
- (rx : zrange)
- (x8 : expr ℤ)
- (rshiftr ry : zrange)
- (y : expr ℤ)
- (offset : Z) =>
- if
- (s7 =? 2 ^ Z.log2 s7) &&
- (ZRange.normalize ry >>
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftr)%zrange &&
- (ZRange.normalize
- rshiftr <=?
- r[0 ~> s7 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_sub
- (Z.log2 s7)
- (- offset))%expr @
- (#(Z_cast rx)%expr @
- x8,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x8, _) := xv in
- x8) args3
- (v (Compile.reflect x3))
- args2 args0
- (v0
- (Compile.reflect x7))
- (let (x8, _) := xv0 in
- x8);
- Some (Base x8));
+ x10);
+ Some (Base x10));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
+ | (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _
+ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _
+ _ _ _ _ _)%expr_pat => None
| _ => None
end
- | @expr.App _ _ _ s6 _ ($_)%expr _ | @expr.App _ _
- _ s6 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
- _ s6 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s6 _
+ | @expr.App _ _ _ s7 _ ($_)%expr _ | @expr.App _ _
+ _ s7 _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _
+ _ s7 _ (_ @ _)%expr_pat _ | @expr.App _ _ _ s7 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
| _ => None
end;;
- args <- invert_bind_args idc2 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc0 Raw.ident.Literal;
+ args <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc1 Raw.ident.Literal;
+ args2 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc Raw.ident.Z_sub_get_borrow;
match
pattern.type.unify_extracted ((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s2) -> s3)%ptype
+ (((projT1 args1) -> s3) -> s4)%ptype
with
| Some (_, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
((ℤ -> ℤ) -> ℤ)%ptype
- (((projT1 args1) -> s2) -> s3)%ptype
+ (((projT1 args1) -> s3) -> s4)%ptype
then
xv <- ident.unify pattern.ident.Literal
##(projT2 args1);
- v <- type.try_make_transport_cps s2 ℤ;
- v0 <- type.try_make_transport_cps s3 ℤ;
- fv <- (x5 <- (let
+ v <- type.try_make_transport_cps s3 ℤ;
+ v0 <- type.try_make_transport_cps s4 ℤ;
+ fv <- (x6 <- (let
'(r1, r2)%zrange := range in
- fun (s4 : Z) (rx : zrange)
- (x5 : expr ℤ) (ry : zrange)
- (y : expr ℤ) =>
+ fun (rs : zrange) (s5 : Z)
+ (rx : zrange) (x6 : expr ℤ)
+ (ry : zrange) (y : expr ℤ) =>
if
- (s4 =? 2 ^ Z.log2 s4) &&
+ (s5 =? 2 ^ Z.log2 s5) &&
(ZRange.normalize ry <=?
- r[0 ~> s4 - 1])%zrange
+ r[0 ~> s5 - 1])%zrange &&
+ is_bounded_by_bool s5
+ (ZRange.normalize rs)
then
Some
(#(Z_cast2 (r1, r2))%expr @
- (#(fancy_sub (Z.log2 s4) 0)%expr @
- (#(Z_cast rx)%expr @ x5,
+ (#(fancy_sub (Z.log2 s5) 0)%expr @
+ (#(Z_cast rx)%expr @ x6,
#(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x5, _) := xv in x5) args0
- (v (Compile.reflect x3)) args
- (v0 (Compile.reflect x4));
- Some (Base x5));
+ else None) args2
+ (let (x6, _) := xv in x6) args0
+ (v (Compile.reflect x4)) args
+ (v0 (Compile.reflect x5));
+ Some (Base x6));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | @expr.App _ _ _ s3 _ ($_)%expr _ | @expr.App _ _ _ s3 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s3 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s3 _
+ | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s4 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
- | @expr.App _ _ _ s2 _ ($_)%expr _ | @expr.App _ _ _ s2 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s2 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s2 _
+ | @expr.App _ _ _ s3 _ ($_)%expr _ | @expr.App _ _ _ s3 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s3 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s3 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
+ | (@expr.Ident _ _ _ t0 idc0 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end
| (@expr.Ident _ _ _ t idc @ x3 @ x2 @ x1 @ x0)%expr_pat =>
match x3 with
- | @expr.Ident _ _ _ t0 idc0 =>
+ | (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat =>
match x2 with
- | @expr.Ident _ _ _ t1 idc1 =>
+ | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5 =>
match x1 with
- | @expr.Ident _ _ _ t2 idc2 =>
- match x0 with
- | @expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t3 idc3)
- x4 =>
- match x4 with
- | (@expr.Ident _ _ _ t4 idc4 @ x6 @ x5)%expr_pat =>
- match x6 with
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- @expr.Ident _ _ _ t8 idc8))%expr_pat =>
- match x5 with
- | @expr.Ident _ _ _ t9 idc9 =>
- args <- invert_bind_args idc9
- Raw.ident.Literal;
- args0 <- invert_bind_args idc8
- Raw.ident.Literal;
- args1 <- invert_bind_args idc7
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc6
- Raw.ident.Z_land;
- args3 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc3
- Raw.ident.Z_cast;
- args6 <- invert_bind_args idc2
- Raw.ident.Literal;
- args7 <- invert_bind_args idc1
- Raw.ident.Literal;
- args8 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> (projT1 args7)) ->
- (projT1 args6)) ->
- (s9 -> (projT1 args0)) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, (_, _, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) ->
- (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) ->
- (projT1 args7)) ->
- (projT1 args6)) ->
- (s9 -> (projT1 args0)) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args8);
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args7);
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args6);
- v <- type.try_make_transport_cps s9 ℤ;
- xv2 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args0);
- xv3 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x11 <- (let
- '(r1, r2)%zrange :=
- range in
- fun
- (s10 cc xx : Z)
- (rshiftl rland
- ry : zrange)
- (y : expr ℤ)
- (mask
- offset : Z) =>
- if
- (s10 =?
- 2 ^ Z.log2 s10) &&
- (ZRange.normalize
- rland <<
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftl)%zrange &&
- (ZRange.normalize
- rshiftl <=?
- r[0 ~> s10 - 1])%zrange &&
- (ZRange.normalize
- ry &'
- ZRange.normalize
- (ZRange.constant
- mask) <=?
- ZRange.normalize
- rland)%zrange &&
- (mask =?
- Z.ones
- (Z.log2 s10 -
- offset)) &&
- (0 <=? offset) &&
- (offset <=?
- Z.log2 s10)
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_addc
- (Z.log2
- s10)
- offset)%expr @
- ((##cc)%expr,
- (##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x11, _) :=
- xv in
- x11)
- (let (x11, _) :=
- xv0 in
- x11)
- (let (x11, _) :=
- xv1 in
- x11) args5 args3
- args1
- (v
- (Compile.reflect
- x10))
- (let (x11, _) :=
- xv2 in
- x11)
- (let (x11, _) :=
- xv3 in
- x11);
- Some (Base x11));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- @expr.Abs _ _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- (_ @ _)))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ ($_)%expr _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (_ @ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _
- _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _
- _ _ _ _ _ @ _))%expr_pat => None
- | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | _ => None
- end;;
- match x6 with
- | @expr.App _ _ _ s6 _
- (@expr.Ident _ _ _ t5 idc5) x7 =>
- match x5 with
- | @expr.Ident _ _ _ t6 idc6 =>
- args <- invert_bind_args idc6
- Raw.ident.Literal;
- args0 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc3
- Raw.ident.Z_cast;
- args3 <- invert_bind_args idc2
- Raw.ident.Literal;
- args4 <- invert_bind_args idc1
- Raw.ident.Literal;
- args5 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) -> (projT1 args4)) ->
- (projT1 args3)) ->
- s6 -> (projT1 args))%ptype
- with
- | Some (_, _, _, (_, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) ->
- (projT1 args4)) ->
- (projT1 args3)) ->
- s6 -> (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args5);
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args4);
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args3);
- v <- type.try_make_transport_cps s6 ℤ;
- xv2 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x8 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s7 cc xx : Z)
- (rshiftr
- ry : zrange)
- (y : expr ℤ)
- (offset : Z) =>
- if
- (s7 =?
- 2 ^ Z.log2 s7) &&
- (ZRange.normalize
- ry >>
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftr)%zrange &&
- (ZRange.normalize
- rshiftr <=?
- r[0 ~> s7 - 1])%zrange
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_addc
- (Z.log2 s7)
- (- offset))%expr @
- ((##cc)%expr,
- (##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x8, _) :=
- xv in
- x8)
- (let (x8, _) :=
- xv0 in
- x8)
- (let (x8, _) :=
- xv1 in
- x8) args2 args0
- (v
- (Compile.reflect
- x7))
- (let (x8, _) :=
- xv2 in
- x8);
- Some (Base x8));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | @expr.App _ _ _ s6 _ ($_)%expr _ | @expr.App
- _ _ _ s6 _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s6 _ (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s6 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | _ => None
- end;;
- args <- invert_bind_args idc3 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc2 Raw.ident.Literal;
- args1 <- invert_bind_args idc1 Raw.ident.Literal;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) ->
- (projT1 args0)) -> s3)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) ->
- (projT1 args0)) -> s3)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s3 ℤ;
- fv <- (x5 <- (let
- '(r1, r2)%zrange := range in
- fun (s4 cc xx : Z)
- (ry : zrange) (y : expr ℤ) =>
- if
- (s4 =? 2 ^ Z.log2 s4) &&
- (ZRange.normalize ry <=?
- r[0 ~> s4 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_addc (Z.log2 s4) 0)%expr @
- ((##cc)%expr, (##xx)%expr,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x5, _) := xv in x5)
- (let (x5, _) := xv0 in x5)
- (let (x5, _) := xv1 in x5) args
- (v (Compile.reflect x4));
- Some (Base x5));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s3 _ ($_)%expr _ | @expr.App _ _ _ s3
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s3 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s3 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t2 idc2) x4 =>
- match x0 with
- | @expr.Ident _ _ _ t3 idc3 =>
- match x4 with
- | (@expr.Ident _ _ _ t4 idc4 @ x6 @ x5)%expr_pat =>
- match x6 with
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- @expr.Ident _ _ _ t8 idc8))%expr_pat =>
- match x5 with
- | @expr.Ident _ _ _ t9 idc9 =>
- args <- invert_bind_args idc9
- Raw.ident.Literal;
- args0 <- invert_bind_args idc8
- Raw.ident.Literal;
- args1 <- invert_bind_args idc7
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc6
- Raw.ident.Z_land;
- args3 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc3
- Raw.ident.Literal;
- args6 <- invert_bind_args idc2
- Raw.ident.Z_cast;
- args7 <- invert_bind_args idc1
- Raw.ident.Literal;
- args8 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> (projT1 args7)) ->
- (s9 -> (projT1 args0)) ->
- (projT1 args)) -> (projT1 args5))%ptype
- with
- | Some (_, _, (_, _, _), _)%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ) ->
- ℤ)%ptype
- ((((projT1 args8) ->
- (projT1 args7)) ->
- (s9 -> (projT1 args0)) ->
- (projT1 args)) ->
- (projT1 args5))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args8);
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args7);
- v <- type.try_make_transport_cps s9 ℤ;
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args0);
- xv2 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- xv3 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args5);
- fv <- (x11 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s10 cc : Z)
- (rshiftl rland
- ry : zrange)
- (y : expr ℤ)
- (mask offset
- xx : Z) =>
- if
- (s10 =?
- 2 ^ Z.log2 s10) &&
- (ZRange.normalize
- rland <<
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftl)%zrange &&
- (ZRange.normalize
- rshiftl <=?
- r[0 ~> s10 - 1])%zrange &&
- (ZRange.normalize
- ry &'
- ZRange.normalize
- (ZRange.constant
- mask) <=?
- ZRange.normalize
- rland)%zrange &&
- (mask =?
- Z.ones
- (Z.log2 s10 -
- offset)) &&
- (0 <=? offset) &&
- (offset <=?
- Z.log2 s10)
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_addc
- (Z.log2
- s10)
- offset)%expr @
- ((##cc)%expr,
- (##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x11, _) :=
- xv in
- x11)
- (let (x11, _) :=
- xv0 in
- x11) args6 args3
- args1
- (v
- (Compile.reflect
- x10))
- (let (x11, _) :=
- xv1 in
- x11)
- (let (x11, _) :=
- xv2 in
- x11)
- (let (x11, _) :=
- xv3 in
- x11);
- Some (Base x11));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- @expr.Abs _ _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- (_ @ _)))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ ($_)%expr _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (_ @ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _
- _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _
- _ _ _ _ _ @ _))%expr_pat => None
- | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | _ => None
- end;;
- match x6 with
- | @expr.App _ _ _ s6 _
- (@expr.Ident _ _ _ t5 idc5) x7 =>
- match x5 with
- | @expr.Ident _ _ _ t6 idc6 =>
- args <- invert_bind_args idc6
- Raw.ident.Literal;
- args0 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc3
- Raw.ident.Literal;
- args3 <- invert_bind_args idc2
- Raw.ident.Z_cast;
- args4 <- invert_bind_args idc1
- Raw.ident.Literal;
- args5 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args5) -> (projT1 args4)) ->
- s6 -> (projT1 args)) ->
- (projT1 args2))%ptype
- with
- | Some (_, _, (_, _), _)%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args5) ->
- (projT1 args4)) ->
- s6 -> (projT1 args)) ->
- (projT1 args2))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args5);
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args4);
- v <- type.try_make_transport_cps s6 ℤ;
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- xv2 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args2);
- fv <- (x8 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s7 cc : Z)
- (rshiftr
- ry : zrange)
- (y : expr ℤ)
- (offset xx : Z)
- =>
- if
- (s7 =?
- 2 ^ Z.log2 s7) &&
- (ZRange.normalize
- ry >>
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftr)%zrange &&
- (ZRange.normalize
- rshiftr <=?
- r[0 ~> s7 - 1])%zrange
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_addc
- (Z.log2 s7)
- (- offset))%expr @
- ((##cc)%expr,
- (##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x8, _) :=
- xv in
- x8)
- (let (x8, _) :=
- xv0 in
- x8) args3 args0
- (v
- (Compile.reflect
- x7))
- (let (x8, _) :=
- xv1 in
- x8)
- (let (x8, _) :=
- xv2 in
- x8);
- Some (Base x8));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | @expr.App _ _ _ s6 _ ($_)%expr _ | @expr.App
- _ _ _ s6 _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s6 _ (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s6 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | _ => None
- end;;
- args <- invert_bind_args idc3 Raw.ident.Literal;
- args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc1 Raw.ident.Literal;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) -> s3) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) -> s3) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s3 ℤ;
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x5 <- (let
- '(r1, r2)%zrange := range in
- fun (s4 cc : Z) (rx : zrange)
- (x5 : expr ℤ) (yy : Z) =>
- if
- (s4 =? 2 ^ Z.log2 s4) &&
- (ZRange.normalize
- (ZRange.constant yy) <=?
- r[0 ~> s4 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_addc (Z.log2 s4) 0)%expr @
- ((##cc)%expr,
- #(Z_cast rx)%expr @ x5,
- (##yy)%expr)))%expr_pat
- else None)
- (let (x5, _) := xv in x5)
- (let (x5, _) := xv0 in x5)
- args0 (v (Compile.reflect x4))
- (let (x5, _) := xv1 in x5);
- Some (Base x5));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t3 idc3)
- x5 =>
- match x5 with
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8)) @ @expr.Ident _ _
- _ t9 idc9)%expr_pat =>
- args <- invert_bind_args idc9 Raw.ident.Literal;
- args0 <- invert_bind_args idc8
- Raw.ident.Literal;
- args1 <- invert_bind_args idc7 Raw.ident.Z_cast;
- _ <- invert_bind_args idc6 Raw.ident.Z_land;
- args3 <- invert_bind_args idc5 Raw.ident.Z_cast;
- _ <- invert_bind_args idc4 Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args6 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args7 <- invert_bind_args idc1
- Raw.ident.Literal;
- args8 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> (projT1 args7)) -> s3) ->
- (s10 -> (projT1 args0)) -> (projT1 args))%ptype
- with
- | Some (_, _, _, (_, _, _))%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> (projT1 args7)) ->
- s3) ->
- (s10 -> (projT1 args0)) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args8);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args7);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s10 ℤ;
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- xv2 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x12 <- (let
- '(r1, r2)%zrange := range
- in
- fun (s11 cc : Z)
- (rx : zrange)
- (x12 : expr ℤ)
- (rshiftl rland
- ry : zrange)
- (y : expr ℤ)
- (mask offset : Z) =>
- if
- (s11 =? 2 ^ Z.log2 s11) &&
- (ZRange.normalize rland <<
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize rshiftl)%zrange &&
- (ZRange.normalize rshiftl <=?
- r[0 ~> s11 - 1])%zrange &&
- (ZRange.normalize ry &'
- ZRange.normalize
- (ZRange.constant mask) <=?
- ZRange.normalize rland)%zrange &&
- (mask =?
- Z.ones
- (Z.log2 s11 - offset)) &&
- (0 <=? offset) &&
- (offset <=? Z.log2 s11)
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_addc
- (Z.log2 s11)
- offset)%expr @
- ((##cc)%expr,
- #(Z_cast rx)%expr @
- x12,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x12, _) := xv in x12)
- (let (x12, _) := xv0 in
- x12) args6
- (v (Compile.reflect x4))
- args5 args3 args1
- (v0 (Compile.reflect x11))
- (let (x12, _) := xv1 in
- x12)
- (let (x12, _) := xv2 in
- x12);
- Some (Base x12));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8)) @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8)) @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8)) @ (_ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8)) @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- ($_)%expr)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Abs _ _ _ _ _ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- (_ @ _))) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.LetIn _ _ _ _ _ _ _)) @ _)%expr_pat =>
- None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ ($_)%expr _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Abs _ _ _ _ _ _) _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (_ @ _) _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)) @ _)%expr_pat =>
- None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _)) @
- _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _
- _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _
- _ _ _ _ @ _)) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ #(_)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _
- _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _)) @
- _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _)) @
- _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _)) @
- _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _
- _ _ _) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @ #(_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _
- _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Abs _ _ _ _ _ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _
- _ _ _ @ _)%expr_pat => None
- | _ => None
- end;;
- match x4 with
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8)) @ @expr.Ident _ _
- _ t9 idc9)%expr_pat =>
- args <- invert_bind_args idc9 Raw.ident.Literal;
- args0 <- invert_bind_args idc8
- Raw.ident.Literal;
- args1 <- invert_bind_args idc7 Raw.ident.Z_cast;
- _ <- invert_bind_args idc6 Raw.ident.Z_land;
- args3 <- invert_bind_args idc5 Raw.ident.Z_cast;
- _ <- invert_bind_args idc4 Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args6 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args7 <- invert_bind_args idc1
- Raw.ident.Literal;
- args8 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> (projT1 args7)) ->
- (s10 -> (projT1 args0)) -> (projT1 args)) ->
- s4)%ptype
- with
- | Some (_, _, (_, _, _), _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> (projT1 args7)) ->
- (s10 -> (projT1 args0)) ->
- (projT1 args)) -> s4)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args8);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args7);
- v <- type.try_make_transport_cps s10 ℤ;
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- xv2 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v0 <- type.try_make_transport_cps s4 ℤ;
- fv <- (x12 <- (let
- '(r1, r2)%zrange := range
- in
- fun (s11 cc : Z)
- (rshiftl rland
- ry : zrange)
- (y : expr ℤ)
- (mask offset : Z)
- (rx : zrange)
- (x12 : expr ℤ) =>
- if
- (s11 =? 2 ^ Z.log2 s11) &&
- (ZRange.normalize rland <<
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize rshiftl)%zrange &&
- (ZRange.normalize rshiftl <=?
- r[0 ~> s11 - 1])%zrange &&
- (ZRange.normalize ry &'
- ZRange.normalize
- (ZRange.constant mask) <=?
- ZRange.normalize rland)%zrange &&
- (mask =?
- Z.ones
- (Z.log2 s11 - offset)) &&
- (0 <=? offset) &&
- (offset <=? Z.log2 s11)
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_addc
- (Z.log2 s11)
- offset)%expr @
- ((##cc)%expr,
- #(Z_cast rx)%expr @
- x12,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x12, _) := xv in x12)
- (let (x12, _) := xv0 in
- x12) args6 args3 args1
- (v (Compile.reflect x11))
- (let (x12, _) := xv1 in
- x12)
- (let (x12, _) := xv2 in
- x12) args5
- (v0 (Compile.reflect x5));
- Some (Base x12));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8)) @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8)) @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8)) @ (_ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8)) @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- ($_)%expr)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Abs _ _ _ _ _ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- (_ @ _))) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.LetIn _ _ _ _ _ _ _)) @ _)%expr_pat =>
- None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ ($_)%expr _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Abs _ _ _ _ _ _) _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (_ @ _) _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)) @ _)%expr_pat =>
- None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _)) @
- _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _
- _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _
- _ _ _ _ @ _)) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ #(_)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _
- _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _)) @
- _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _)) @
- _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _)) @
- _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _
- _ _ _) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @ #(_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _
- _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Abs _ _ _ _ _ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _
- _ _ _ @ _)%expr_pat => None
- | _ => None
- end;;
- match x5 with
- | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 @ @expr.Ident _ _
- _ t6 idc6)%expr_pat =>
- args <- invert_bind_args idc6 Raw.ident.Literal;
- args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
- _ <- invert_bind_args idc4 Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args4 <- invert_bind_args idc1
- Raw.ident.Literal;
- args5 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) -> (projT1 args4)) -> s3) ->
- s7 -> (projT1 args))%ptype
- with
- | Some (_, _, _, (_, _))%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) -> (projT1 args4)) ->
- s3) -> s7 -> (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args5);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args4);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s7 ℤ;
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x9 <- (let
- '(r1, r2)%zrange := range in
- fun (s8 cc : Z)
- (rx : zrange)
- (x9 : expr ℤ)
- (rshiftr ry : zrange)
- (y : expr ℤ) (offset : Z)
- =>
- if
- (s8 =? 2 ^ Z.log2 s8) &&
- (ZRange.normalize ry >>
- ZRange.normalize
- (ZRange.constant offset) <=?
- ZRange.normalize rshiftr)%zrange &&
- (ZRange.normalize rshiftr <=?
- r[0 ~> s8 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_addc
- (Z.log2 s8)
- (- offset))%expr @
- ((##cc)%expr,
- #(Z_cast rx)%expr @ x9,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x9, _) := xv in x9)
- (let (x9, _) := xv0 in x9)
- args3
- (v (Compile.reflect x4))
- args2 args0
- (v0 (Compile.reflect x8))
- (let (x9, _) := xv1 in x9);
- Some (Base x9));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 @ (_ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- ($_)%expr _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (_ @ _) _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- None
- | (@expr.Ident _ _ _ t4 idc4 @ #(_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _
- _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _
- _ _ _ @ _)%expr_pat => None
- | _ => None
- end;;
- match x4 with
- | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 @ @expr.Ident _ _
- _ t6 idc6)%expr_pat =>
- args <- invert_bind_args idc6 Raw.ident.Literal;
- args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
- _ <- invert_bind_args idc4 Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args4 <- invert_bind_args idc1
- Raw.ident.Literal;
- args5 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args5) -> (projT1 args4)) ->
- s7 -> (projT1 args)) -> s4)%ptype
- with
- | Some (_, _, (_, _), _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args5) -> (projT1 args4)) ->
- s7 -> (projT1 args)) -> s4)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args5);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args4);
- v <- type.try_make_transport_cps s7 ℤ;
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- v0 <- type.try_make_transport_cps s4 ℤ;
- fv <- (x9 <- (let
- '(r1, r2)%zrange := range in
- fun (s8 cc : Z)
- (rshiftr ry : zrange)
- (y : expr ℤ) (offset : Z)
- (rx : zrange)
- (x9 : expr ℤ) =>
- if
- (s8 =? 2 ^ Z.log2 s8) &&
- (ZRange.normalize ry >>
- ZRange.normalize
- (ZRange.constant offset) <=?
- ZRange.normalize rshiftr)%zrange &&
- (ZRange.normalize rshiftr <=?
- r[0 ~> s8 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_addc
- (Z.log2 s8)
- (- offset))%expr @
- ((##cc)%expr,
- #(Z_cast rx)%expr @ x9,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x9, _) := xv in x9)
- (let (x9, _) := xv0 in x9)
- args3 args0
- (v (Compile.reflect x8))
- (let (x9, _) := xv1 in x9)
- args2
- (v0 (Compile.reflect x5));
- Some (Base x9));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 @ (_ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- ($_)%expr _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (_ @ _) _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s7 _
- (@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
- None
- | (@expr.Ident _ _ _ t4 idc4 @ #(_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _
- _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _
- _ _ _ @ _)%expr_pat => None
- | _ => None
- end;;
- args <- invert_bind_args idc3 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc1 Raw.ident.Literal;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) -> s3) ->
- s4)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) -> s3) ->
- s4)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s4 ℤ;
- fv <- (x6 <- (let
- '(r1, r2)%zrange := range in
- fun (s5 cc : Z) (rx : zrange)
- (x6 : expr ℤ) (ry : zrange)
- (y : expr ℤ) =>
- if
- (s5 =? 2 ^ Z.log2 s5) &&
- (ZRange.normalize ry <=?
- r[0 ~> s5 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_addc (Z.log2 s5) 0)%expr @
- ((##cc)%expr,
- #(Z_cast rx)%expr @ x6,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x6, _) := xv in x6)
- (let (x6, _) := xv0 in x6)
- args0 (v (Compile.reflect x4))
- args (v0 (Compile.reflect x5));
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s4 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s3 _ ($_)%expr _ | @expr.App _ _ _ s3 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s3 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s3 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4 =>
- match x1 with
- | @expr.Ident _ _ _ t2 idc2 =>
- match x0 with
- | @expr.Ident _ _ _ t3 idc3 =>
- args <- invert_bind_args idc3 Raw.ident.Literal;
- args0 <- invert_bind_args idc2 Raw.ident.Literal;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> (projT1 args0)) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> (projT1 args0)) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s3 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x5 <- (let
- '(r1, r2)%zrange := range in
- fun (s4 : Z) (rc : zrange)
- (c : expr ℤ) (xx yy : Z) =>
- if
- (s4 =? 2 ^ Z.log2 s4) &&
- (ZRange.normalize
- (ZRange.constant yy) <=?
- r[0 ~> s4 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_addc (Z.log2 s4) 0)%expr @
- (#(Z_cast rc)%expr @ c,
- (##xx)%expr, (##yy)%expr)))%expr_pat
- else None)
- (let (x5, _) := xv in x5) args1
- (v (Compile.reflect x4))
- (let (x5, _) := xv0 in x5)
- (let (x5, _) := xv1 in x5);
- Some (Base x5));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t3 idc3)
- x5 =>
- match x5 with
- | (@expr.Ident _ _ _ t4 idc4 @ x7 @ x6)%expr_pat =>
- match x7 with
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8))%expr_pat =>
- match x6 with
- | @expr.Ident _ _ _ t9 idc9 =>
- args <- invert_bind_args idc9
- Raw.ident.Literal;
- args0 <- invert_bind_args idc8
- Raw.ident.Literal;
- args1 <- invert_bind_args idc7
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc6
- Raw.ident.Z_land;
- args3 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc3
- Raw.ident.Z_cast;
- args6 <- invert_bind_args idc2
- Raw.ident.Literal;
- args7 <- invert_bind_args idc1
- Raw.ident.Z_cast;
- args8 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> s3) ->
- (projT1 args6)) ->
- (s10 -> (projT1 args0)) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, (_, _, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) ->
- (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> s3) ->
- (projT1 args6)) ->
- (s10 -> (projT1 args0)) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args8);
- v <- type.try_make_transport_cps s3 ℤ;
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args6);
- v0 <- type.try_make_transport_cps s10 ℤ;
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args0);
- xv2 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x12 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s11 : Z)
- (rc : zrange)
- (c : expr ℤ)
- (xx : Z)
- (rshiftl rland
- ry : zrange)
- (y : expr ℤ)
- (mask
- offset : Z) =>
- if
- (s11 =?
- 2 ^ Z.log2 s11) &&
- (ZRange.normalize
- rland <<
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftl)%zrange &&
- (ZRange.normalize
- rshiftl <=?
- r[0 ~> s11 - 1])%zrange &&
- (ZRange.normalize
- ry &'
- ZRange.normalize
- (ZRange.constant
- mask) <=?
- ZRange.normalize
- rland)%zrange &&
- (mask =?
- Z.ones
- (Z.log2 s11 -
- offset)) &&
- (0 <=? offset) &&
- (offset <=?
- Z.log2 s11)
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_addc
- (Z.log2
- s11)
- offset)%expr @
- (#(Z_cast rc)%expr @
- c,
- (##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x12, _) :=
- xv in
- x12) args7
- (v
- (Compile.reflect
- x4))
- (let (x12, _) :=
- xv0 in
- x12) args5 args3
- args1
- (v0
- (Compile.reflect
- x11))
- (let (x12, _) :=
- xv1 in
- x12)
- (let (x12, _) :=
- xv2 in
- x12);
- Some (Base x12));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Abs _ _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- (_ @ _)))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ ($_)%expr _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (_ @ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _
- _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _
- _ _ _ _ _ @ _))%expr_pat => None
- | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | _ => None
- end;;
- match x7 with
- | @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 =>
- match x6 with
- | @expr.Ident _ _ _ t6 idc6 =>
- args <- invert_bind_args idc6
- Raw.ident.Literal;
- args0 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc3
- Raw.ident.Z_cast;
- args3 <- invert_bind_args idc2
- Raw.ident.Literal;
- args4 <- invert_bind_args idc1
- Raw.ident.Z_cast;
- args5 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) -> s3) ->
- (projT1 args3)) ->
- s7 -> (projT1 args))%ptype
- with
- | Some (_, _, _, (_, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) -> s3) ->
- (projT1 args3)) ->
- s7 -> (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args5);
- v <- type.try_make_transport_cps s3 ℤ;
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args3);
- v0 <- type.try_make_transport_cps s7 ℤ;
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x9 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s8 : Z)
- (rc : zrange)
- (c : expr ℤ)
- (xx : Z)
- (rshiftr
- ry : zrange)
- (y : expr ℤ)
- (offset : Z) =>
- if
- (s8 =?
- 2 ^ Z.log2 s8) &&
- (ZRange.normalize
- ry >>
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftr)%zrange &&
- (ZRange.normalize
- rshiftr <=?
- r[0 ~> s8 - 1])%zrange
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_addc
- (Z.log2 s8)
- (- offset))%expr @
- (#(Z_cast rc)%expr @
- c,
- (##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x9, _) :=
- xv in
- x9) args4
- (v
- (Compile.reflect
- x4))
- (let (x9, _) :=
- xv0 in
- x9) args2 args0
- (v0
- (Compile.reflect
- x8))
- (let (x9, _) :=
- xv1 in
- x9);
- Some (Base x9));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | @expr.App _ _ _ s7 _ ($_)%expr _ | @expr.App
- _ _ _ s7 _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s7 _ (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s7 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | _ => None
- end;;
- args <- invert_bind_args idc3 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc2 Raw.ident.Literal;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> (projT1 args0)) ->
- s4)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> (projT1 args0)) ->
- s4)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s3 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v0 <- type.try_make_transport_cps s4 ℤ;
- fv <- (x6 <- (let
- '(r1, r2)%zrange := range in
- fun (s5 : Z) (rc : zrange)
- (c : expr ℤ) (xx : Z)
- (ry : zrange) (y : expr ℤ) =>
- if
- (s5 =? 2 ^ Z.log2 s5) &&
- (ZRange.normalize ry <=?
- r[0 ~> s5 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_addc (Z.log2 s5) 0)%expr @
- (#(Z_cast rc)%expr @ c,
- (##xx)%expr,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x6, _) := xv in x6) args1
- (v (Compile.reflect x4))
- (let (x6, _) := xv0 in x6) args
- (v0 (Compile.reflect x5));
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s4 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5 =>
+ | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6 =>
match x0 with
- | @expr.Ident _ _ _ t3 idc3 =>
- match x5 with
- | (@expr.Ident _ _ _ t4 idc4 @ x7 @ x6)%expr_pat =>
- match x7 with
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8))%expr_pat =>
- match x6 with
- | @expr.Ident _ _ _ t9 idc9 =>
- args <- invert_bind_args idc9
- Raw.ident.Literal;
- args0 <- invert_bind_args idc8
- Raw.ident.Literal;
- args1 <- invert_bind_args idc7
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc6
- Raw.ident.Z_land;
- args3 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc3
- Raw.ident.Literal;
- args6 <- invert_bind_args idc2
- Raw.ident.Z_cast;
- args7 <- invert_bind_args idc1
- Raw.ident.Z_cast;
- args8 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> s3) ->
- (s10 -> (projT1 args0)) ->
- (projT1 args)) -> (projT1 args5))%ptype
- with
- | Some (_, _, (_, _, _), _)%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ) ->
- ℤ)%ptype
- ((((projT1 args8) -> s3) ->
- (s10 -> (projT1 args0)) ->
- (projT1 args)) ->
- (projT1 args5))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args8);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s10 ℤ;
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args0);
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- xv2 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args5);
- fv <- (x12 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s11 : Z)
- (rc : zrange)
- (c : expr ℤ)
- (rshiftl rland
- ry : zrange)
- (y : expr ℤ)
- (mask offset
- xx : Z) =>
- if
- (s11 =?
- 2 ^ Z.log2 s11) &&
- (ZRange.normalize
- rland <<
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftl)%zrange &&
- (ZRange.normalize
- rshiftl <=?
- r[0 ~> s11 - 1])%zrange &&
- (ZRange.normalize
- ry &'
- ZRange.normalize
- (ZRange.constant
- mask) <=?
- ZRange.normalize
- rland)%zrange &&
- (mask =?
- Z.ones
- (Z.log2 s11 -
- offset)) &&
- (0 <=? offset) &&
- (offset <=?
- Z.log2 s11)
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_addc
- (Z.log2
- s11)
- offset)%expr @
- (#(Z_cast rc)%expr @
- c,
- (##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x12, _) :=
- xv in
- x12) args7
- (v
- (Compile.reflect
- x4)) args6
- args3 args1
- (v0
- (Compile.reflect
- x11))
- (let (x12, _) :=
- xv0 in
- x12)
- (let (x12, _) :=
- xv1 in
- x12)
- (let (x12, _) :=
- xv2 in
- x12);
- Some (Base x12));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Abs _ _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- (_ @ _)))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ ($_)%expr _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (_ @ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _
- _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _
- _ _ _ _ _ @ _))%expr_pat => None
- | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | _ => None
- end;;
- match x7 with
- | @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 =>
- match x6 with
- | @expr.Ident _ _ _ t6 idc6 =>
- args <- invert_bind_args idc6
- Raw.ident.Literal;
- args0 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc3
- Raw.ident.Literal;
- args3 <- invert_bind_args idc2
- Raw.ident.Z_cast;
- args4 <- invert_bind_args idc1
- Raw.ident.Z_cast;
- args5 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args5) -> s3) ->
- s7 -> (projT1 args)) ->
- (projT1 args2))%ptype
- with
- | Some (_, _, (_, _), _)%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args5) -> s3) ->
- s7 -> (projT1 args)) ->
- (projT1 args2))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args5);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s7 ℤ;
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args2);
- fv <- (x9 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s8 : Z)
- (rc : zrange)
- (c : expr ℤ)
- (rshiftr
- ry : zrange)
- (y : expr ℤ)
- (offset xx : Z)
- =>
- if
- (s8 =?
- 2 ^ Z.log2 s8) &&
- (ZRange.normalize
- ry >>
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftr)%zrange &&
- (ZRange.normalize
- rshiftr <=?
- r[0 ~> s8 - 1])%zrange
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_addc
- (Z.log2 s8)
- (- offset))%expr @
- (#(Z_cast rc)%expr @
- c,
- (##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x9, _) :=
- xv in
- x9) args4
- (v
- (Compile.reflect
- x4)) args3
- args0
- (v0
- (Compile.reflect
- x8))
- (let (x9, _) :=
- xv0 in
- x9)
- (let (x9, _) :=
- xv1 in
- x9);
- Some (Base x9));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | @expr.App _ _ _ s7 _ ($_)%expr _ | @expr.App
- _ _ _ s7 _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s7 _ (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s7 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | _ => None
- end;;
- args <- invert_bind_args idc3 Raw.ident.Literal;
- args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_add_with_get_carry;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> s4) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> s4) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s4 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x6 <- (let
- '(r1, r2)%zrange := range in
- fun (s5 : Z) (rc : zrange)
- (c : expr ℤ) (rx : zrange)
- (x6 : expr ℤ) (yy : Z) =>
- if
- (s5 =? 2 ^ Z.log2 s5) &&
- (ZRange.normalize
- (ZRange.constant yy) <=?
- r[0 ~> s5 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_addc (Z.log2 s5) 0)%expr @
- (#(Z_cast rc)%expr @ c,
- #(Z_cast rx)%expr @ x6,
- (##yy)%expr)))%expr_pat
- else None)
- (let (x6, _) := xv in x6) args1
- (v (Compile.reflect x4)) args0
- (v0 (Compile.reflect x5))
- (let (x6, _) := xv0 in x6);
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3)
- x6 =>
- match x6 with
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- @expr.Ident _ _ _ t8 idc8)) @ @expr.Ident _ _
- _ t9 idc9)%expr_pat =>
- args <- invert_bind_args idc9 Raw.ident.Literal;
- args0 <- invert_bind_args idc8
- Raw.ident.Literal;
- args1 <- invert_bind_args idc7 Raw.ident.Z_cast;
- _ <- invert_bind_args idc6 Raw.ident.Z_land;
- args3 <- invert_bind_args idc5 Raw.ident.Z_cast;
- _ <- invert_bind_args idc4 Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args6 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args7 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args8 <- invert_bind_args idc0
+ | @expr.App _ _ _ s6 _ (@expr.Ident _ _ _ t4 idc4)
+ x7 =>
+ match x7 with
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @
+ (@expr.Ident _ _ _ t11 idc11 @ @expr.Ident _ _ _
+ t12 idc12))%expr_pat =>
+ args <- invert_bind_args idc12
+ Raw.ident.Literal;
+ args0 <- invert_bind_args idc11
+ Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc10
Raw.ident.Literal;
+ args2 <- invert_bind_args idc9 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc8 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc7 Raw.ident.Z_land;
+ args5 <- invert_bind_args idc6 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc5 Raw.ident.Z_shiftl;
+ args7 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ args8 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args9 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args10 <- invert_bind_args idc1
+ Raw.ident.Literal;
+ args11 <- invert_bind_args idc0
+ Raw.ident.Z_cast;
_ <- invert_bind_args idc
Raw.ident.Z_add_with_get_carry;
match
pattern.type.unify_extracted
(((ℤ -> ℤ) -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> s3) -> s4) ->
- (s11 -> (projT1 args0)) -> (projT1 args))%ptype
+ ((((projT1 args10) -> s4) -> s5) ->
+ (s12 -> (projT1 args1)) -> (projT1 args))%ptype
with
| Some (_, _, _, (_, _, _))%zrange =>
if
type.type_beq base.type base.type.type_beq
(((ℤ -> ℤ) -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> s3) -> s4) ->
- (s11 -> (projT1 args0)) ->
+ ((((projT1 args10) -> s4) -> s5) ->
+ (s12 -> (projT1 args1)) ->
(projT1 args))%ptype
then
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args8);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s4 ℤ;
- v1 <- type.try_make_transport_cps s11 ℤ;
+ ##(projT2 args10);
+ v <- type.try_make_transport_cps s4 ℤ;
+ v0 <- type.try_make_transport_cps s5 ℤ;
+ v1 <- type.try_make_transport_cps s12 ℤ;
xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
+ ##(projT2 args1);
xv1 <- ident.unify pattern.ident.Literal
##(projT2 args);
- fv <- (x13 <- (let
+ fv <- (x16 <- (let
'(r1, r2)%zrange := range
in
- fun (s12 : Z)
- (rc : zrange)
+ fun (rs : zrange)
+ (s15 : Z) (rc : zrange)
(c : expr ℤ)
(rx : zrange)
- (x13 : expr ℤ)
+ (x16 : expr ℤ)
(rshiftl rland
ry : zrange)
(y : expr ℤ)
- (mask offset : Z) =>
+ (rmask : zrange)
+ (mask : Z)
+ (roffset : zrange)
+ (offset : Z) =>
if
- (s12 =? 2 ^ Z.log2 s12) &&
+ (s15 =? 2 ^ Z.log2 s15) &&
(ZRange.normalize rland <<
ZRange.normalize
(ZRange.constant
@@ -7434,2039 +4726,1031 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
(ZRange.constant mask) <=?
ZRange.normalize rland)%zrange &&
(ZRange.normalize rshiftl <=?
- r[0 ~> s12 - 1])%zrange &&
+ r[0 ~> s15 - 1])%zrange &&
(mask =?
Z.ones
- (Z.log2 s12 - offset)) &&
+ (Z.log2 s15 - offset)) &&
(0 <=? offset) &&
- (offset <=? Z.log2 s12)
+ (offset <=? Z.log2 s15) &&
+ is_bounded_by_bool s15
+ (ZRange.normalize rs) &&
+ is_bounded_by_bool mask
+ (ZRange.normalize rmask) &&
+ is_bounded_by_bool offset
+ (ZRange.normalize
+ roffset)
then
Some
(#(Z_cast2 (r1, r2))%expr @
(#(fancy_addc
- (Z.log2 s12)
+ (Z.log2 s15)
offset)%expr @
(#(Z_cast rc)%expr @
c,
#(Z_cast rx)%expr @
- x13,
+ x16,
#(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x13, _) := xv in x13)
- args7
- (v (Compile.reflect x4))
- args6
- (v0 (Compile.reflect x5))
- args5 args3 args1
- (v1 (Compile.reflect x12))
- (let (x13, _) := xv0 in
- x13)
- (let (x13, _) := xv1 in
- x13);
- Some (Base x13));
+ else None) args11
+ (let (x16, _) := xv in x16)
+ args9
+ (v (Compile.reflect x5))
+ args8
+ (v0 (Compile.reflect x6))
+ args7 args5 args3
+ (v1 (Compile.reflect x13))
+ args2
+ (let (x16, _) := xv0 in
+ x16) args0
+ (let (x16, _) := xv1 in
+ x16);
+ Some (Base x16));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- @expr.Ident _ _ _ t8 idc8)) @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- @expr.Ident _ _ _ t8 idc8)) @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- @expr.Ident _ _ _ t8 idc8)) @ (_ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- @expr.Ident _ _ _ t8 idc8)) @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @
+ (@expr.Ident _ _ _ t11 idc11 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @
+ (@expr.Ident _ _ _ t11 idc11 @ @expr.Abs _ _ _ _
+ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @
+ (@expr.Ident _ _ _ t11 idc11 @ (_ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @
+ (@expr.Ident _ _ _ t11 idc11 @ @expr.LetIn _ _ _
+ _ _ _ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ (_ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ ($_)%expr))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Abs _ _ _ _
+ _ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ (_ @ _)))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.LetIn _ _ _
+ _ _ _ _))) @ _)%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @ #
+ (_))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
($_)%expr)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
@expr.Abs _ _ _ _ _ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- (_ @ _))) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (($_)%expr @ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Abs _ _ _ _ _ _ @ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (_ @ _ @ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
@expr.LetIn _ _ _ _ _ _ _)) @ _)%expr_pat =>
None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ ($_)%expr _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Abs _ _ _ _ _ _) _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (_ @ _) _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)) @ _)%expr_pat =>
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ ($_)%expr _ @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Abs _ _ _ _ _ _) _ @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (_ @ _) _ @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)) @ _)%expr_pat =>
None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _)) @
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ #(_) @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ ($_)%expr @ _)) @
_)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Abs _ _ _ _
_ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.LetIn _ _ _
_ _ _ _ @ _)) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ #(_)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ #(_)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _ _
_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _)) @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (#(_) @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (($_)%expr @ _)) @
_)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
(@expr.Abs _ _ _ _ _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _)) @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (($_)%expr @ _ @ _)) @
_)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
(@expr.Abs _ _ _ _ _ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _)) @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (_ @ _ @ _ @ _)) @
_)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
(@expr.LetIn _ _ _ _ _ _ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
(@expr.LetIn _ _ _ _ _ _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _ _
_ _ _) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @ #(_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _
+ | (@expr.Ident _ _ _ t5 idc5 @ #(_) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _
_ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.Abs _ _ _ _ _ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.LetIn _ _ _ _ _ _ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _
_ _ _ @ _)%expr_pat => None
| _ => None
end;;
- match x5 with
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- @expr.Ident _ _ _ t8 idc8)) @ @expr.Ident _ _
- _ t9 idc9)%expr_pat =>
- args <- invert_bind_args idc9 Raw.ident.Literal;
- args0 <- invert_bind_args idc8
- Raw.ident.Literal;
- args1 <- invert_bind_args idc7 Raw.ident.Z_cast;
- _ <- invert_bind_args idc6 Raw.ident.Z_land;
- args3 <- invert_bind_args idc5 Raw.ident.Z_cast;
- _ <- invert_bind_args idc4 Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args6 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args7 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args8 <- invert_bind_args idc0
+ match x6 with
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @
+ (@expr.Ident _ _ _ t11 idc11 @ @expr.Ident _ _ _
+ t12 idc12))%expr_pat =>
+ args <- invert_bind_args idc12
+ Raw.ident.Literal;
+ args0 <- invert_bind_args idc11
+ Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc10
Raw.ident.Literal;
+ args2 <- invert_bind_args idc9 Raw.ident.Z_cast;
+ args3 <- invert_bind_args idc8 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc7 Raw.ident.Z_land;
+ args5 <- invert_bind_args idc6 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc5 Raw.ident.Z_shiftl;
+ args7 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ args8 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args9 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args10 <- invert_bind_args idc1
+ Raw.ident.Literal;
+ args11 <- invert_bind_args idc0
+ Raw.ident.Z_cast;
_ <- invert_bind_args idc
Raw.ident.Z_add_with_get_carry;
match
pattern.type.unify_extracted
(((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> s3) ->
- (s11 -> (projT1 args0)) -> (projT1 args)) ->
- s5)%ptype
+ ((((projT1 args10) -> s4) ->
+ (s12 -> (projT1 args1)) -> (projT1 args)) ->
+ s6)%ptype
with
| Some (_, _, (_, _, _), _)%zrange =>
if
type.type_beq base.type base.type.type_beq
(((ℤ -> ℤ) -> (ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> s3) ->
- (s11 -> (projT1 args0)) ->
- (projT1 args)) -> s5)%ptype
+ ((((projT1 args10) -> s4) ->
+ (s12 -> (projT1 args1)) ->
+ (projT1 args)) -> s6)%ptype
then
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args8);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s11 ℤ;
+ ##(projT2 args10);
+ v <- type.try_make_transport_cps s4 ℤ;
+ v0 <- type.try_make_transport_cps s12 ℤ;
xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
+ ##(projT2 args1);
xv1 <- ident.unify pattern.ident.Literal
##(projT2 args);
- v1 <- type.try_make_transport_cps s5 ℤ;
- fv <- (x13 <- (let
+ v1 <- type.try_make_transport_cps s6 ℤ;
+ fv <- (x16 <- (let
'(r1, r2)%zrange := range
in
- fun (s12 : Z)
- (rc : zrange)
+ fun (rs : zrange)
+ (s15 : Z) (rc : zrange)
(c : expr ℤ)
(rshiftl rland
ry : zrange)
(y : expr ℤ)
- (mask offset : Z)
+ (rmask : zrange)
+ (mask : Z)
+ (roffset : zrange)
+ (offset : Z)
(rx : zrange)
- (x13 : expr ℤ) =>
+ (x16 : expr ℤ) =>
if
- (s12 =? 2 ^ Z.log2 s12) &&
+ (s15 =? 2 ^ Z.log2 s15) &&
(ZRange.normalize rland <<
ZRange.normalize
(ZRange.constant
offset) <=?
ZRange.normalize rshiftl)%zrange &&
(ZRange.normalize rshiftl <=?
- r[0 ~> s12 - 1])%zrange &&
+ r[0 ~> s15 - 1])%zrange &&
(ZRange.normalize ry &'
ZRange.normalize
(ZRange.constant mask) <=?
ZRange.normalize rland)%zrange &&
(mask =?
Z.ones
- (Z.log2 s12 - offset)) &&
+ (Z.log2 s15 - offset)) &&
(0 <=? offset) &&
- (offset <=? Z.log2 s12)
+ (offset <=? Z.log2 s15) &&
+ is_bounded_by_bool s15
+ (ZRange.normalize rs) &&
+ is_bounded_by_bool mask
+ (ZRange.normalize rmask) &&
+ is_bounded_by_bool offset
+ (ZRange.normalize
+ roffset)
then
Some
(#(Z_cast2 (r1, r2))%expr @
(#(fancy_addc
- (Z.log2 s12)
+ (Z.log2 s15)
offset)%expr @
(#(Z_cast rc)%expr @
c,
#(Z_cast rx)%expr @
- x13,
+ x16,
#(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x13, _) := xv in x13)
- args7
- (v (Compile.reflect x4))
- args6 args3 args1
- (v0 (Compile.reflect x12))
- (let (x13, _) := xv0 in
- x13)
- (let (x13, _) := xv1 in
- x13) args5
- (v1 (Compile.reflect x6));
- Some (Base x13));
+ else None) args11
+ (let (x16, _) := xv in x16)
+ args9
+ (v (Compile.reflect x5))
+ args8 args5 args3
+ (v0 (Compile.reflect x13))
+ args2
+ (let (x16, _) := xv0 in
+ x16) args0
+ (let (x16, _) := xv1 in
+ x16) args7
+ (v1 (Compile.reflect x7));
+ Some (Base x16));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- @expr.Ident _ _ _ t8 idc8)) @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- @expr.Ident _ _ _ t8 idc8)) @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- @expr.Ident _ _ _ t8 idc8)) @ (_ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- @expr.Ident _ _ _ t8 idc8)) @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @
+ (@expr.Ident _ _ _ t11 idc11 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @
+ (@expr.Ident _ _ _ t11 idc11 @ @expr.Abs _ _ _ _
+ _ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @
+ (@expr.Ident _ _ _ t11 idc11 @ (_ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @
+ (@expr.Ident _ _ _ t11 idc11 @ @expr.LetIn _ _ _
+ _ _ _ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ (_ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _ _ _
+ t10 idc10))) @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ ($_)%expr))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Abs _ _ _ _
+ _ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ (_ @ _)))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.LetIn _ _ _
+ _ _ _ _))) @ _)%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @ #
+ (_))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
($_)%expr)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
@expr.Abs _ _ _ _ _ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- (_ @ _))) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (($_)%expr @ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Abs _ _ _ _ _ _ @ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (_ @ _ @ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
@expr.LetIn _ _ _ _ _ _ _)) @ _)%expr_pat =>
None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ ($_)%expr _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Abs _ _ _ _ _ _) _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (_ @ _) _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)) @ _)%expr_pat =>
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ ($_)%expr _ @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Abs _ _ _ _ _ _) _ @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (_ @ _) _ @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _)) @ _)%expr_pat =>
None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _)) @
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ #(_) @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ ($_)%expr @ _)) @
_)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Abs _ _ _ _
_ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.LetIn _ _ _
_ _ _ _ @ _)) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ #(_)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _
+ | (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ #(_)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _ _ _
_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _)) @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (#(_) @ _)) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (($_)%expr @ _)) @
_)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
(@expr.Abs _ _ _ _ _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _ @ _)) @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (($_)%expr @ _ @ _)) @
_)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
(@expr.Abs _ _ _ _ _ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _)) @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (_ @ _ @ _ @ _)) @
_)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
(@expr.LetIn _ _ _ _ _ _ _ @ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
(@expr.LetIn _ _ _ _ _ _ _ @ _)) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _
+ (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _ _ _
_ _ _) @ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @ #(_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _
+ | (@expr.Ident _ _ _ t5 idc5 @ #(_) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _
_ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (($_)%expr @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.Abs _ _ _ _ _ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ (_ @ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @
+ (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @
(@expr.LetIn _ _ _ _ _ _ _ @ _) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _
_ _ _ @ _)%expr_pat => None
| _ => None
end;;
- match x6 with
- | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
- (@expr.Ident _ _ _ t5 idc5) x9 @ @expr.Ident _ _
- _ t6 idc6)%expr_pat =>
- args <- invert_bind_args idc6 Raw.ident.Literal;
- args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
- _ <- invert_bind_args idc4 Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args4 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args5 <- invert_bind_args idc0
+ match x7 with
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _
+ t8 idc8))%expr_pat =>
+ args <- invert_bind_args idc8 Raw.ident.Literal;
+ args0 <- invert_bind_args idc7 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc6 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc5 Raw.ident.Z_shiftr;
+ args3 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args5 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args6 <- invert_bind_args idc1
Raw.ident.Literal;
+ args7 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc
Raw.ident.Z_add_with_get_carry;
match
pattern.type.unify_extracted
(((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) -> s3) -> s4) ->
- s8 -> (projT1 args))%ptype
+ ((((projT1 args6) -> s4) -> s5) ->
+ s9 -> (projT1 args))%ptype
with
| Some (_, _, _, (_, _))%zrange =>
if
type.type_beq base.type base.type.type_beq
(((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) -> s3) -> s4) ->
- s8 -> (projT1 args))%ptype
+ ((((projT1 args6) -> s4) -> s5) ->
+ s9 -> (projT1 args))%ptype
then
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args5);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s4 ℤ;
- v1 <- type.try_make_transport_cps s8 ℤ;
+ ##(projT2 args6);
+ v <- type.try_make_transport_cps s4 ℤ;
+ v0 <- type.try_make_transport_cps s5 ℤ;
+ v1 <- type.try_make_transport_cps s9 ℤ;
xv0 <- ident.unify pattern.ident.Literal
##(projT2 args);
- fv <- (x10 <- (let
+ fv <- (x12 <- (let
'(r1, r2)%zrange := range
in
- fun (s9 : Z) (rc : zrange)
+ fun (rs : zrange)
+ (s11 : Z) (rc : zrange)
(c : expr ℤ)
(rx : zrange)
- (x10 : expr ℤ)
+ (x12 : expr ℤ)
(rshiftr ry : zrange)
(y : expr ℤ)
+ (roffset : zrange)
(offset : Z) =>
if
- (s9 =? 2 ^ Z.log2 s9) &&
+ (s11 =? 2 ^ Z.log2 s11) &&
(ZRange.normalize ry >>
ZRange.normalize
(ZRange.constant
offset) <=?
ZRange.normalize rshiftr)%zrange &&
(ZRange.normalize rshiftr <=?
- r[0 ~> s9 - 1])%zrange
+ r[0 ~> s11 - 1])%zrange &&
+ is_bounded_by_bool s11
+ (ZRange.normalize rs) &&
+ is_bounded_by_bool offset
+ (ZRange.normalize
+ roffset)
then
Some
(#(Z_cast2 (r1, r2))%expr @
(#(fancy_addc
- (Z.log2 s9)
+ (Z.log2 s11)
(- offset))%expr @
(#(Z_cast rc)%expr @
c,
#(Z_cast rx)%expr @
- x10,
+ x12,
#(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x10, _) := xv in x10)
+ else None) args7
+ (let (x12, _) := xv in x12)
+ args5
+ (v (Compile.reflect x5))
args4
- (v (Compile.reflect x4))
- args3
- (v0 (Compile.reflect x5))
- args2 args0
- (v1 (Compile.reflect x9))
- (let (x10, _) := xv0 in
- x10);
- Some (Base x10));
+ (v0 (Compile.reflect x6))
+ args3 args1
+ (v1 (Compile.reflect x10))
+ args0
+ (let (x12, _) := xv0 in
+ x12);
+ Some (Base x12));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
- (@expr.Ident _ _ _ t5 idc5) x9 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
- (@expr.Ident _ _ _ t5 idc5) x9 @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
- (@expr.Ident _ _ _ t5 idc5) x9 @ (_ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
- (@expr.Ident _ _ _ t5 idc5) x9 @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.Ident _ _ _ t7 idc7 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Abs _ _ _ _ _
+ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.Ident _ _ _ t7 idc7 @ (_ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.LetIn _ _ _ _
+ _ _ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Abs _ _
+ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @ (_ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.LetIn _
+ _ _ _ _ _ _)%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
($_)%expr _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
(@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
(_ @ _) _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
(@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
None
- | (@expr.Ident _ _ _ t4 idc4 @ #(_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _
+ | (@expr.Ident _ _ _ t5 idc5 @ #(_) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _
_ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _
_ _ _ @ _)%expr_pat => None
| _ => None
end;;
- match x5 with
- | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
- (@expr.Ident _ _ _ t5 idc5) x9 @ @expr.Ident _ _
- _ t6 idc6)%expr_pat =>
- args <- invert_bind_args idc6 Raw.ident.Literal;
- args0 <- invert_bind_args idc5 Raw.ident.Z_cast;
- _ <- invert_bind_args idc4 Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc3 Raw.ident.Z_cast;
- args3 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args4 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args5 <- invert_bind_args idc0
+ match x6 with
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident _ _ _
+ t8 idc8))%expr_pat =>
+ args <- invert_bind_args idc8 Raw.ident.Literal;
+ args0 <- invert_bind_args idc7 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc6 Raw.ident.Z_cast;
+ _ <- invert_bind_args idc5 Raw.ident.Z_shiftr;
+ args3 <- invert_bind_args idc4 Raw.ident.Z_cast;
+ args4 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args5 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args6 <- invert_bind_args idc1
Raw.ident.Literal;
+ args7 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc
Raw.ident.Z_add_with_get_carry;
match
pattern.type.unify_extracted
(((ℤ -> ℤ) -> ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args5) -> s3) ->
- s8 -> (projT1 args)) -> s5)%ptype
+ ((((projT1 args6) -> s4) ->
+ s9 -> (projT1 args)) -> s6)%ptype
with
| Some (_, _, (_, _), _)%zrange =>
if
type.type_beq base.type base.type.type_beq
(((ℤ -> ℤ) -> ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args5) -> s3) ->
- s8 -> (projT1 args)) -> s5)%ptype
+ ((((projT1 args6) -> s4) ->
+ s9 -> (projT1 args)) -> s6)%ptype
then
xv <- ident.unify pattern.ident.Literal
- ##(projT2 args5);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s8 ℤ;
+ ##(projT2 args6);
+ v <- type.try_make_transport_cps s4 ℤ;
+ v0 <- type.try_make_transport_cps s9 ℤ;
xv0 <- ident.unify pattern.ident.Literal
##(projT2 args);
- v1 <- type.try_make_transport_cps s5 ℤ;
- fv <- (x10 <- (let
+ v1 <- type.try_make_transport_cps s6 ℤ;
+ fv <- (x12 <- (let
'(r1, r2)%zrange := range
in
- fun (s9 : Z) (rc : zrange)
+ fun (rs : zrange)
+ (s11 : Z) (rc : zrange)
(c : expr ℤ)
(rshiftr ry : zrange)
(y : expr ℤ)
+ (roffset : zrange)
(offset : Z)
(rx : zrange)
- (x10 : expr ℤ) =>
+ (x12 : expr ℤ) =>
if
- (s9 =? 2 ^ Z.log2 s9) &&
+ (s11 =? 2 ^ Z.log2 s11) &&
(ZRange.normalize ry >>
ZRange.normalize
(ZRange.constant
offset) <=?
ZRange.normalize rshiftr)%zrange &&
(ZRange.normalize rshiftr <=?
- r[0 ~> s9 - 1])%zrange
+ r[0 ~> s11 - 1])%zrange &&
+ is_bounded_by_bool s11
+ (ZRange.normalize rs) &&
+ is_bounded_by_bool offset
+ (ZRange.normalize
+ roffset)
then
Some
(#(Z_cast2 (r1, r2))%expr @
(#(fancy_addc
- (Z.log2 s9)
+ (Z.log2 s11)
(- offset))%expr @
(#(Z_cast rc)%expr @
c,
#(Z_cast rx)%expr @
- x10,
+ x12,
#(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x10, _) := xv in x10)
- args4
- (v (Compile.reflect x4))
- args3 args0
- (v0 (Compile.reflect x9))
- (let (x10, _) := xv0 in
- x10) args2
- (v1 (Compile.reflect x6));
- Some (Base x10));
+ else None) args7
+ (let (x12, _) := xv in x12)
+ args5
+ (v (Compile.reflect x5))
+ args4 args1
+ (v0 (Compile.reflect x10))
+ args0
+ (let (x12, _) := xv0 in
+ x12) args3
+ (v1 (Compile.reflect x7));
+ Some (Base x12));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
- (@expr.Ident _ _ _ t5 idc5) x9 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
- (@expr.Ident _ _ _ t5 idc5) x9 @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
- (@expr.Ident _ _ _ t5 idc5) x9 @ (_ @ _))%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
- (@expr.Ident _ _ _ t5 idc5) x9 @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.Ident _ _ _ t7 idc7 @ ($_)%expr))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Abs _ _ _ _ _
+ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.Ident _ _ _ t7 idc7 @ (_ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.LetIn _ _ _ _
+ _ _ _))%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.Abs _ _
+ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @ (_ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 @ @expr.LetIn _
+ _ _ _ _ _ _)%expr_pat => None
+ | (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
($_)%expr _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
(@expr.Abs _ _ _ _ _ _) _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
(_ @ _) _ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.App _ _ _ s8 _
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.App _ _ _ s9 _
(@expr.LetIn _ _ _ _ _ _ _) _ @ _)%expr_pat =>
None
- | (@expr.Ident _ _ _ t4 idc4 @ #(_) @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ ($_)%expr @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.Abs _ _ _ _ _
+ | (@expr.Ident _ _ _ t5 idc5 @ #(_) @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr @ _)%expr_pat |
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _ _ _
_ @ _)%expr_pat |
- (@expr.Ident _ _ _ t4 idc4 @ @expr.LetIn _ _ _ _
+ (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _ _ _
_ _ _ @ _)%expr_pat => None
| _ => None
end;;
- args <- invert_bind_args idc3 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
+ args <- invert_bind_args idc4 Raw.ident.Z_cast;
+ args0 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc1 Raw.ident.Literal;
+ args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc
Raw.ident.Z_add_with_get_carry;
match
pattern.type.unify_extracted
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> s4) -> s5)%ptype
+ ((((projT1 args2) -> s4) -> s5) -> s6)%ptype
with
| Some (_, _, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> s4) -> s5)%ptype
+ ((((projT1 args2) -> s4) -> s5) -> s6)%ptype
then
xv <- ident.unify pattern.ident.Literal
##(projT2 args2);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s4 ℤ;
- v1 <- type.try_make_transport_cps s5 ℤ;
- fv <- (x7 <- (let
+ v <- type.try_make_transport_cps s4 ℤ;
+ v0 <- type.try_make_transport_cps s5 ℤ;
+ v1 <- type.try_make_transport_cps s6 ℤ;
+ fv <- (x8 <- (let
'(r1, r2)%zrange := range in
- fun (s6 : Z) (rc : zrange)
- (c : expr ℤ) (rx : zrange)
- (x7 : expr ℤ) (ry : zrange)
- (y : expr ℤ) =>
+ fun (rs : zrange) (s7 : Z)
+ (rc : zrange) (c : expr ℤ)
+ (rx : zrange) (x8 : expr ℤ)
+ (ry : zrange) (y : expr ℤ) =>
if
- (s6 =? 2 ^ Z.log2 s6) &&
+ (s7 =? 2 ^ Z.log2 s7) &&
(ZRange.normalize ry <=?
- r[0 ~> s6 - 1])%zrange
+ r[0 ~> s7 - 1])%zrange &&
+ is_bounded_by_bool s7
+ (ZRange.normalize rs)
then
Some
(#(Z_cast2 (r1, r2))%expr @
- (#(fancy_addc (Z.log2 s6) 0)%expr @
+ (#(fancy_addc (Z.log2 s7) 0)%expr @
(#(Z_cast rc)%expr @ c,
- #(Z_cast rx)%expr @ x7,
+ #(Z_cast rx)%expr @ x8,
#(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x7, _) := xv in x7) args1
- (v (Compile.reflect x4)) args0
- (v0 (Compile.reflect x5)) args
- (v1 (Compile.reflect x6));
- Some (Base x7));
+ else None) args3
+ (let (x8, _) := xv in x8) args1
+ (v (Compile.reflect x5)) args0
+ (v0 (Compile.reflect x6)) args
+ (v1 (Compile.reflect x7));
+ Some (Base x8));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _ _ s5
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s5 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s5 _
+ | @expr.App _ _ _ s6 _ ($_)%expr _ | @expr.App _ _ _ s6
+ _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s6 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s6 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
- | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s4 _
+ | @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _ _ s5 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s5 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s5 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
- | @expr.App _ _ _ s3 _ ($_)%expr _ | @expr.App _ _ _ s3 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s3 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s3 _
+ | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s4 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
+ | (@expr.Ident _ _ _ t0 idc0 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end;;
match x3 with
- | @expr.Ident _ _ _ t0 idc0 =>
+ | (@expr.Ident _ _ _ t0 idc0 @ @expr.Ident _ _ _ t1 idc1)%expr_pat =>
match x2 with
- | @expr.Ident _ _ _ t1 idc1 =>
- match x1 with
- | @expr.Ident _ _ _ t2 idc2 =>
- match x0 with
- | @expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t3 idc3)
- x4 =>
- match x4 with
- | (@expr.Ident _ _ _ t4 idc4 @ x6 @ x5)%expr_pat =>
- match x6 with
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- @expr.Ident _ _ _ t8 idc8))%expr_pat =>
- match x5 with
- | @expr.Ident _ _ _ t9 idc9 =>
- args <- invert_bind_args idc9
- Raw.ident.Literal;
- args0 <- invert_bind_args idc8
- Raw.ident.Literal;
- args1 <- invert_bind_args idc7
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc6
- Raw.ident.Z_land;
- args3 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc3
- Raw.ident.Z_cast;
- args6 <- invert_bind_args idc2
- Raw.ident.Literal;
- args7 <- invert_bind_args idc1
- Raw.ident.Literal;
- args8 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_with_get_borrow;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> (projT1 args7)) ->
- (projT1 args6)) ->
- (s9 -> (projT1 args0)) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, (_, _, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) ->
- (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) ->
- (projT1 args7)) ->
- (projT1 args6)) ->
- (s9 -> (projT1 args0)) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args8);
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args7);
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args6);
- v <- type.try_make_transport_cps s9 ℤ;
- xv2 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args0);
- xv3 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x11 <- (let
- '(r1, r2)%zrange :=
- range in
- fun
- (s10 bb xx : Z)
- (rshiftl rland
- ry : zrange)
- (y : expr ℤ)
- (mask
- offset : Z) =>
- if
- (s10 =?
- 2 ^ Z.log2 s10) &&
- (ZRange.normalize
- rland <<
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftl)%zrange &&
- (ZRange.normalize
- rshiftl <=?
- r[0 ~> s10 - 1])%zrange &&
- (ZRange.normalize
- ry &'
- ZRange.normalize
- (ZRange.constant
- mask) <=?
- ZRange.normalize
- rland)%zrange &&
- (mask =?
- Z.ones
- (Z.log2 s10 -
- offset)) &&
- (0 <=? offset) &&
- (offset <=?
- Z.log2 s10)
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_subb
- (Z.log2
- s10)
- offset)%expr @
- ((##bb)%expr,
- (##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x11, _) :=
- xv in
- x11)
- (let (x11, _) :=
- xv0 in
- x11)
- (let (x11, _) :=
- xv1 in
- x11) args5 args3
- args1
- (v
- (Compile.reflect
- x10))
- (let (x11, _) :=
- xv2 in
- x11)
- (let (x11, _) :=
- xv3 in
- x11);
- Some (Base x11));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- @expr.Abs _ _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- (_ @ _)))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Ident _ _ _ t7 idc7) x10 @
- @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ ($_)%expr _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (_ @ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s9 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _
- _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _
- _ _ _ _ _ @ _))%expr_pat => None
- | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | _ => None
- end;;
- match x6 with
- | @expr.App _ _ _ s6 _
- (@expr.Ident _ _ _ t5 idc5) x7 =>
- match x5 with
- | @expr.Ident _ _ _ t6 idc6 =>
- args <- invert_bind_args idc6
- Raw.ident.Literal;
- args0 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc3
- Raw.ident.Z_cast;
- args3 <- invert_bind_args idc2
- Raw.ident.Literal;
- args4 <- invert_bind_args idc1
- Raw.ident.Literal;
- args5 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_with_get_borrow;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) -> (projT1 args4)) ->
- (projT1 args3)) ->
- s6 -> (projT1 args))%ptype
- with
- | Some (_, _, _, (_, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) ->
- (projT1 args4)) ->
- (projT1 args3)) ->
- s6 -> (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args5);
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args4);
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args3);
- v <- type.try_make_transport_cps s6 ℤ;
- xv2 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x8 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s7 bb xx : Z)
- (rshiftr
- ry : zrange)
- (y : expr ℤ)
- (offset : Z) =>
- if
- (s7 =?
- 2 ^ Z.log2 s7) &&
- (ZRange.normalize
- ry >>
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftr)%zrange &&
- (ZRange.normalize
- rshiftr <=?
- r[0 ~> s7 - 1])%zrange
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_subb
- (Z.log2 s7)
- (- offset))%expr @
- ((##bb)%expr,
- (##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x8, _) :=
- xv in
- x8)
- (let (x8, _) :=
- xv0 in
- x8)
- (let (x8, _) :=
- xv1 in
- x8) args2 args0
- (v
- (Compile.reflect
- x7))
- (let (x8, _) :=
- xv2 in
- x8);
- Some (Base x8));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | @expr.App _ _ _ s6 _ ($_)%expr _ | @expr.App
- _ _ _ s6 _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s6 _ (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s6 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | _ => None
- end;;
- args <- invert_bind_args idc3 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc2 Raw.ident.Literal;
- args1 <- invert_bind_args idc1 Raw.ident.Literal;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_with_get_borrow;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) ->
- (projT1 args0)) -> s3)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) ->
- (projT1 args0)) -> s3)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v <- type.try_make_transport_cps s3 ℤ;
- fv <- (x5 <- (let
- '(r1, r2)%zrange := range in
- fun (s4 bb xx : Z)
- (ry : zrange) (y : expr ℤ) =>
- if
- (s4 =? 2 ^ Z.log2 s4) &&
- (ZRange.normalize ry <=?
- r[0 ~> s4 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_subb (Z.log2 s4) 0)%expr @
- ((##bb)%expr, (##xx)%expr,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x5, _) := xv in x5)
- (let (x5, _) := xv0 in x5)
- (let (x5, _) := xv1 in x5) args
- (v (Compile.reflect x4));
- Some (Base x5));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s3 _ ($_)%expr _ | @expr.App _ _ _ s3
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s3 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s3 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t2 idc2) x4 =>
- match x0 with
- | @expr.Ident _ _ _ t3 idc3 =>
- args <- invert_bind_args idc3 Raw.ident.Literal;
- args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc1 Raw.ident.Literal;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_with_get_borrow;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) -> s3) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) -> s3) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s3 ℤ;
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x5 <- (let
- '(r1, r2)%zrange := range in
- fun (s4 bb : Z) (rx : zrange)
- (x5 : expr ℤ) (yy : Z) =>
- if
- (s4 =? 2 ^ Z.log2 s4) &&
- (ZRange.normalize
- (ZRange.constant yy) <=?
- r[0 ~> s4 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_subb (Z.log2 s4) 0)%expr @
- ((##bb)%expr,
- #(Z_cast rx)%expr @ x5,
- (##yy)%expr)))%expr_pat
- else None)
- (let (x5, _) := xv in x5)
- (let (x5, _) := xv0 in x5)
- args0 (v (Compile.reflect x4))
- (let (x5, _) := xv1 in x5);
- Some (Base x5));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t3 idc3)
- x5 =>
- match x5 with
- | (@expr.Ident _ _ _ t4 idc4 @ x7 @ x6)%expr_pat =>
- match x7 with
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8))%expr_pat =>
- match x6 with
- | @expr.Ident _ _ _ t9 idc9 =>
- args <- invert_bind_args idc9
- Raw.ident.Literal;
- args0 <- invert_bind_args idc8
- Raw.ident.Literal;
- args1 <- invert_bind_args idc7
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc6
- Raw.ident.Z_land;
- args3 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc3
- Raw.ident.Z_cast;
- args6 <- invert_bind_args idc2
- Raw.ident.Z_cast;
- args7 <- invert_bind_args idc1
- Raw.ident.Literal;
- args8 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_with_get_borrow;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> (projT1 args7)) ->
- s3) ->
- (s10 -> (projT1 args0)) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, (_, _, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) ->
- (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) ->
- (projT1 args7)) -> s3) ->
- (s10 -> (projT1 args0)) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args8);
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args7);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s10 ℤ;
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args0);
- xv2 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x12 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s11 bb : Z)
- (rx : zrange)
- (x12 : expr ℤ)
- (rshiftl rland
- ry : zrange)
- (y : expr ℤ)
- (mask
- offset : Z) =>
- if
- (s11 =?
- 2 ^ Z.log2 s11) &&
- (ZRange.normalize
- rland <<
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftl)%zrange &&
- (ZRange.normalize
- rshiftl <=?
- r[0 ~> s11 - 1])%zrange &&
- (ZRange.normalize
- ry &'
- ZRange.normalize
- (ZRange.constant
- mask) <=?
- ZRange.normalize
- rland)%zrange &&
- (mask =?
- Z.ones
- (Z.log2 s11 -
- offset)) &&
- (0 <=? offset) &&
- (offset <=?
- Z.log2 s11)
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_subb
- (Z.log2
- s11)
- offset)%expr @
- ((##bb)%expr,
- #(Z_cast rx)%expr @
- x12,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x12, _) :=
- xv in
- x12)
- (let (x12, _) :=
- xv0 in
- x12) args6
- (v
- (Compile.reflect
- x4)) args5
- args3 args1
- (v0
- (Compile.reflect
- x11))
- (let (x12, _) :=
- xv1 in
- x12)
- (let (x12, _) :=
- xv2 in
- x12);
- Some (Base x12));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Abs _ _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- (_ @ _)))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ ($_)%expr _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (_ @ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _
- _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _
- _ _ _ _ _ @ _))%expr_pat => None
- | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | _ => None
- end;;
- match x7 with
- | @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 =>
- match x6 with
- | @expr.Ident _ _ _ t6 idc6 =>
- args <- invert_bind_args idc6
- Raw.ident.Literal;
- args0 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc3
- Raw.ident.Z_cast;
- args3 <- invert_bind_args idc2
- Raw.ident.Z_cast;
- args4 <- invert_bind_args idc1
- Raw.ident.Literal;
- args5 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_with_get_borrow;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) -> (projT1 args4)) ->
- s3) -> s7 -> (projT1 args))%ptype
- with
- | Some (_, _, _, (_, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) ->
- (projT1 args4)) -> s3) ->
- s7 -> (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args5);
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args4);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s7 ℤ;
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x9 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s8 bb : Z)
- (rx : zrange)
- (x9 : expr ℤ)
- (rshiftr
- ry : zrange)
- (y : expr ℤ)
- (offset : Z) =>
- if
- (s8 =?
- 2 ^ Z.log2 s8) &&
- (ZRange.normalize
- ry >>
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftr)%zrange &&
- (ZRange.normalize
- rshiftr <=?
- r[0 ~> s8 - 1])%zrange
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_subb
- (Z.log2 s8)
- (- offset))%expr @
- ((##bb)%expr,
- #(Z_cast rx)%expr @
- x9,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x9, _) :=
- xv in
- x9)
- (let (x9, _) :=
- xv0 in
- x9) args3
- (v
- (Compile.reflect
- x4)) args2
- args0
- (v0
- (Compile.reflect
- x8))
- (let (x9, _) :=
- xv1 in
- x9);
- Some (Base x9));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | @expr.App _ _ _ s7 _ ($_)%expr _ | @expr.App
- _ _ _ s7 _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s7 _ (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s7 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | _ => None
- end;;
- args <- invert_bind_args idc3 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc1 Raw.ident.Literal;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_with_get_borrow;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) -> s3) ->
- s4)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> (projT1 args1)) -> s3) ->
- s4)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args1);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s4 ℤ;
- fv <- (x6 <- (let
- '(r1, r2)%zrange := range in
- fun (s5 bb : Z) (rx : zrange)
- (x6 : expr ℤ) (ry : zrange)
- (y : expr ℤ) =>
- if
- (s5 =? 2 ^ Z.log2 s5) &&
- (ZRange.normalize ry <=?
- r[0 ~> s5 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_subb (Z.log2 s5) 0)%expr @
- ((##bb)%expr,
- #(Z_cast rx)%expr @ x6,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x6, _) := xv in x6)
- (let (x6, _) := xv0 in x6)
- args0 (v (Compile.reflect x4))
- args (v0 (Compile.reflect x5));
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s4 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s3 _ ($_)%expr _ | @expr.App _ _ _ s3 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s3 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s3 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s3 _ (@expr.Ident _ _ _ t1 idc1) x4 =>
+ | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5 =>
match x1 with
- | @expr.Ident _ _ _ t2 idc2 =>
+ | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3) x6 =>
match x0 with
- | @expr.Ident _ _ _ t3 idc3 =>
- args <- invert_bind_args idc3 Raw.ident.Literal;
- args0 <- invert_bind_args idc2 Raw.ident.Literal;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_with_get_borrow;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> (projT1 args0)) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> (projT1 args0)) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s3 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- xv1 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x5 <- (let
- '(r1, r2)%zrange := range in
- fun (s4 : Z) (rb : zrange)
- (b2 : expr ℤ) (xx yy : Z) =>
- if
- (s4 =? 2 ^ Z.log2 s4) &&
- (ZRange.normalize
- (ZRange.constant yy) <=?
- r[0 ~> s4 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_subb (Z.log2 s4) 0)%expr @
- (#(Z_cast rb)%expr @ b2,
- (##xx)%expr, (##yy)%expr)))%expr_pat
- else None)
- (let (x5, _) := xv in x5) args1
- (v (Compile.reflect x4))
- (let (x5, _) := xv0 in x5)
- (let (x5, _) := xv1 in x5);
- Some (Base x5));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t3 idc3)
- x5 =>
- match x5 with
- | (@expr.Ident _ _ _ t4 idc4 @ x7 @ x6)%expr_pat =>
- match x7 with
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Ident _ _ _ t8 idc8))%expr_pat =>
- match x6 with
- | @expr.Ident _ _ _ t9 idc9 =>
- args <- invert_bind_args idc9
- Raw.ident.Literal;
- args0 <- invert_bind_args idc8
- Raw.ident.Literal;
- args1 <- invert_bind_args idc7
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc6
- Raw.ident.Z_land;
- args3 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc3
- Raw.ident.Z_cast;
- args6 <- invert_bind_args idc2
- Raw.ident.Literal;
- args7 <- invert_bind_args idc1
- Raw.ident.Z_cast;
- args8 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_with_get_borrow;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> s3) ->
- (projT1 args6)) ->
- (s10 -> (projT1 args0)) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, (_, _, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) ->
- (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> s3) ->
- (projT1 args6)) ->
- (s10 -> (projT1 args0)) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args8);
- v <- type.try_make_transport_cps s3 ℤ;
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args6);
- v0 <- type.try_make_transport_cps s10 ℤ;
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args0);
- xv2 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x12 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s11 : Z)
- (rb : zrange)
- (b4 : expr ℤ)
- (xx : Z)
- (rshiftl rland
- ry : zrange)
- (y : expr ℤ)
- (mask
- offset : Z) =>
- if
- (s11 =?
- 2 ^ Z.log2 s11) &&
- (ZRange.normalize
- rland <<
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftl)%zrange &&
- (ZRange.normalize
- rshiftl <=?
- r[0 ~> s11 - 1])%zrange &&
- (ZRange.normalize
- ry &'
- ZRange.normalize
- (ZRange.constant
- mask) <=?
- ZRange.normalize
- rland)%zrange &&
- (mask =?
- Z.ones
- (Z.log2 s11 -
- offset)) &&
- (0 <=? offset) &&
- (offset <=?
- Z.log2 s11)
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_subb
- (Z.log2
- s11)
- offset)%expr @
- (#(Z_cast rb)%expr @
- b4,
- (##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x12, _) :=
- xv in
- x12) args7
- (v
- (Compile.reflect
- x4))
- (let (x12, _) :=
- xv0 in
- x12) args5 args3
- args1
- (v0
- (Compile.reflect
- x11))
- (let (x12, _) :=
- xv1 in
- x12)
- (let (x12, _) :=
- xv2 in
- x12);
- Some (Base x12));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- ($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.Abs _ _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- (_ @ _)))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Ident _ _ _ t7 idc7) x11 @
- @expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ ($_)%expr _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (_ @ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s10 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
- None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _
- _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _
- _ _ _ _ _ @ _))%expr_pat => None
- | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _
- _ _ _)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _
- _ _ _ _ _)%expr_pat => None
- | _ => None
- end;;
- match x7 with
- | @expr.App _ _ _ s7 _
- (@expr.Ident _ _ _ t5 idc5) x8 =>
- match x6 with
- | @expr.Ident _ _ _ t6 idc6 =>
- args <- invert_bind_args idc6
+ | @expr.App _ _ _ s6 _ (@expr.Ident _ _ _ t4 idc4)
+ x7 =>
+ match x7 with
+ | (@expr.Ident _ _ _ t5 idc5 @ x9 @ x8)%expr_pat =>
+ match x9 with
+ | (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Ident _
+ _ _ t10 idc10)))%expr_pat =>
+ match x8 with
+ | (@expr.Ident _ _ _ t11 idc11 @
+ @expr.Ident _ _ _ t12 idc12)%expr_pat =>
+ args <- invert_bind_args idc12
Raw.ident.Literal;
- args0 <- invert_bind_args idc5
- Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
- Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc3
+ args0 <- invert_bind_args idc11
Raw.ident.Z_cast;
- args3 <- invert_bind_args idc2
+ args1 <- invert_bind_args idc10
Raw.ident.Literal;
- args4 <- invert_bind_args idc1
+ args2 <- invert_bind_args idc9
Raw.ident.Z_cast;
- args5 <- invert_bind_args idc0
- Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_with_get_borrow;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) -> s3) ->
- (projT1 args3)) ->
- s7 -> (projT1 args))%ptype
- with
- | Some (_, _, _, (_, _))%zrange =>
- if
- type.type_beq base.type
- base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) -> s3) ->
- (projT1 args3)) ->
- s7 -> (projT1 args))%ptype
- then
- xv <- ident.unify
- pattern.ident.Literal
- ##(projT2 args5);
- v <- type.try_make_transport_cps s3 ℤ;
- xv0 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args3);
- v0 <- type.try_make_transport_cps s7 ℤ;
- xv1 <- ident.unify
- pattern.ident.Literal
- ##(projT2 args);
- fv <- (x9 <- (let
- '(r1, r2)%zrange :=
- range in
- fun (s8 : Z)
- (rb : zrange)
- (b3 : expr ℤ)
- (xx : Z)
- (rshiftr
- ry : zrange)
- (y : expr ℤ)
- (offset : Z) =>
- if
- (s8 =?
- 2 ^ Z.log2 s8) &&
- (ZRange.normalize
- ry >>
- ZRange.normalize
- (ZRange.constant
- offset) <=?
- ZRange.normalize
- rshiftr)%zrange &&
- (ZRange.normalize
- rshiftr <=?
- r[0 ~> s8 - 1])%zrange
- then
- Some
- (#(Z_cast2
- (r1, r2))%expr @
- (#(fancy_subb
- (Z.log2 s8)
- (- offset))%expr @
- (#(Z_cast rb)%expr @
- b3,
- (##xx)%expr,
- #(Z_cast ry)%expr @
- y)))%expr_pat
- else None)
- (let (x9, _) :=
- xv in
- x9) args4
- (v
- (Compile.reflect
- x4))
- (let (x9, _) :=
- xv0 in
- x9) args2 args0
- (v0
- (Compile.reflect
- x8))
- (let (x9, _) :=
- xv1 in
- x9);
- Some (Base x9));
- Some
- (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | _ => None
- end
- | @expr.App _ _ _ s7 _ ($_)%expr _ | @expr.App
- _ _ _ s7 _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s7 _ (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s7 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | _ => None
- end;;
- args <- invert_bind_args idc3 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc2 Raw.ident.Literal;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_with_get_borrow;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> (projT1 args0)) ->
- s4)%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> (projT1 args0)) ->
- s4)%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s3 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args0);
- v0 <- type.try_make_transport_cps s4 ℤ;
- fv <- (x6 <- (let
- '(r1, r2)%zrange := range in
- fun (s5 : Z) (rb : zrange)
- (b2 : expr ℤ) (xx : Z)
- (ry : zrange) (y : expr ℤ) =>
- if
- (s5 =? 2 ^ Z.log2 s5) &&
- (ZRange.normalize ry <=?
- r[0 ~> s5 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_subb (Z.log2 s5) 0)%expr @
- (#(Z_cast rb)%expr @ b2,
- (##xx)%expr,
- #(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x6, _) := xv in x6) args1
- (v (Compile.reflect x4))
- (let (x6, _) := xv0 in x6) args
- (v0 (Compile.reflect x5));
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s4 _
- (@expr.LetIn _ _ _ _ _ _ _) _ => None
- | _ => None
- end
- | @expr.App _ _ _ s4 _ (@expr.Ident _ _ _ t2 idc2) x5 =>
- match x0 with
- | @expr.Ident _ _ _ t3 idc3 =>
- args <- invert_bind_args idc3 Raw.ident.Literal;
- args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
- _ <- invert_bind_args idc
- Raw.ident.Z_sub_with_get_borrow;
- match
- pattern.type.unify_extracted
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> s4) ->
- (projT1 args))%ptype
- with
- | Some (_, _, _, _)%zrange =>
- if
- type.type_beq base.type base.type.type_beq
- (((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> s4) ->
- (projT1 args))%ptype
- then
- xv <- ident.unify pattern.ident.Literal
- ##(projT2 args2);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s4 ℤ;
- xv0 <- ident.unify pattern.ident.Literal
- ##(projT2 args);
- fv <- (x6 <- (let
- '(r1, r2)%zrange := range in
- fun (s5 : Z) (rb : zrange)
- (b2 : expr ℤ) (rx : zrange)
- (x6 : expr ℤ) (yy : Z) =>
- if
- (s5 =? 2 ^ Z.log2 s5) &&
- (ZRange.normalize
- (ZRange.constant yy) <=?
- r[0 ~> s5 - 1])%zrange
- then
- Some
- (#(Z_cast2 (r1, r2))%expr @
- (#(fancy_subb (Z.log2 s5) 0)%expr @
- (#(Z_cast rb)%expr @ b2,
- #(Z_cast rx)%expr @ x6,
- (##yy)%expr)))%expr_pat
- else None)
- (let (x6, _) := xv in x6) args1
- (v (Compile.reflect x4)) args0
- (v0 (Compile.reflect x5))
- (let (x6, _) := xv0 in x6);
- Some (Base x6));
- Some (fv0 <-- fv;
- Base fv0)%under_lets
- else None
- | None => None
- end
- | @expr.App _ _ _ s5 _ (@expr.Ident _ _ _ t3 idc3)
- x6 =>
- match x6 with
- | (@expr.Ident _ _ _ t4 idc4 @ x8 @ x7)%expr_pat =>
- match x8 with
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- @expr.Ident _ _ _ t8 idc8))%expr_pat =>
- match x7 with
- | @expr.Ident _ _ _ t9 idc9 =>
- args <- invert_bind_args idc9
- Raw.ident.Literal;
- args0 <- invert_bind_args idc8
- Raw.ident.Literal;
- args1 <- invert_bind_args idc7
+ args3 <- invert_bind_args idc8
Raw.ident.Z_cast;
- _ <- invert_bind_args idc6
+ _ <- invert_bind_args idc7
Raw.ident.Z_land;
- args3 <- invert_bind_args idc5
+ args5 <- invert_bind_args idc6
Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
+ _ <- invert_bind_args idc5
Raw.ident.Z_shiftl;
- args5 <- invert_bind_args idc3
+ args7 <- invert_bind_args idc4
Raw.ident.Z_cast;
- args6 <- invert_bind_args idc2
+ args8 <- invert_bind_args idc3
Raw.ident.Z_cast;
- args7 <- invert_bind_args idc1
+ args9 <- invert_bind_args idc2
Raw.ident.Z_cast;
- args8 <- invert_bind_args idc0
- Raw.ident.Literal;
+ args10 <- invert_bind_args idc1
+ Raw.ident.Literal;
+ args11 <- invert_bind_args idc0
+ Raw.ident.Z_cast;
_ <- invert_bind_args idc
Raw.ident.Z_sub_with_get_borrow;
match
pattern.type.unify_extracted
(((ℤ -> ℤ) -> ℤ) -> (ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> s3) -> s4) ->
- (s11 -> (projT1 args0)) ->
+ ((((projT1 args10) -> s4) -> s5) ->
+ (s12 -> (projT1 args1)) ->
(projT1 args))%ptype
with
| Some (_, _, _, (_, _, _))%zrange =>
@@ -9475,38 +5759,41 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
base.type.type_beq
(((ℤ -> ℤ) -> ℤ) ->
(ℤ -> ℤ) -> ℤ)%ptype
- ((((projT1 args8) -> s3) -> s4) ->
- (s11 -> (projT1 args0)) ->
+ ((((projT1 args10) -> s4) -> s5) ->
+ (s12 -> (projT1 args1)) ->
(projT1 args))%ptype
then
xv <- ident.unify
pattern.ident.Literal
- ##(projT2 args8);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s4 ℤ;
- v1 <- type.try_make_transport_cps s11 ℤ;
+ ##(projT2 args10);
+ v <- type.try_make_transport_cps s4 ℤ;
+ v0 <- type.try_make_transport_cps s5 ℤ;
+ v1 <- type.try_make_transport_cps s12 ℤ;
xv0 <- ident.unify
pattern.ident.Literal
- ##(projT2 args0);
+ ##(projT2 args1);
xv1 <- ident.unify
pattern.ident.Literal
##(projT2 args);
- fv <- (x13 <- (let
+ fv <- (x16 <- (let
'(r1, r2)%zrange :=
range in
- fun (s12 : Z)
+ fun (rs : zrange)
+ (s15 : Z)
(rb : zrange)
(b4 : expr ℤ)
(rx : zrange)
- (x13 : expr ℤ)
+ (x16 : expr ℤ)
(rshiftl rland
ry : zrange)
(y : expr ℤ)
- (mask
- offset : Z) =>
+ (rmask : zrange)
+ (mask : Z)
+ (roffset : zrange)
+ (offset : Z) =>
if
- (s12 =?
- 2 ^ Z.log2 s12) &&
+ (s15 =?
+ 2 ^ Z.log2 s15) &&
(ZRange.normalize
rland <<
ZRange.normalize
@@ -9516,7 +5803,7 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
rshiftl)%zrange &&
(ZRange.normalize
rshiftl <=?
- r[0 ~> s12 - 1])%zrange &&
+ r[0 ~> s15 - 1])%zrange &&
(ZRange.normalize
ry &'
ZRange.normalize
@@ -9526,174 +5813,234 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
rland)%zrange &&
(mask =?
Z.ones
- (Z.log2 s12 -
+ (Z.log2 s15 -
offset)) &&
(0 <=? offset) &&
(offset <=?
- Z.log2 s12)
+ Z.log2 s15) &&
+ is_bounded_by_bool
+ s15
+ (ZRange.normalize
+ rs) &&
+ is_bounded_by_bool
+ mask
+ (ZRange.normalize
+ rmask) &&
+ is_bounded_by_bool
+ offset
+ (ZRange.normalize
+ roffset)
then
Some
(#(Z_cast2
(r1, r2))%expr @
(#(fancy_subb
(Z.log2
- s12)
+ s15)
offset)%expr @
(#(Z_cast rb)%expr @
b4,
#(Z_cast rx)%expr @
- x13,
+ x16,
#(Z_cast ry)%expr @
y)))%expr_pat
- else None)
- (let (x13, _) :=
+ else None) args11
+ (let (x16, _) :=
xv in
- x13) args7
+ x16) args9
(v
(Compile.reflect
- x4)) args6
+ x5)) args8
(v0
(Compile.reflect
- x5)) args5
- args3 args1
+ x6)) args7
+ args5 args3
(v1
(Compile.reflect
- x12))
- (let (x13, _) :=
+ x13)) args2
+ (let (x16, _) :=
xv0 in
- x13)
- (let (x13, _) :=
+ x16) args0
+ (let (x16, _) :=
xv1 in
- x13);
- Some (Base x13));
+ x16);
+ Some (Base x16));
Some
(fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
+ | (@expr.Ident _ _ _ t11 idc11 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t11 idc11 @ @expr.Abs
+ _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t11 idc11 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t11 idc11 @
+ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
+ | (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ ($_)%expr)))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.Abs _ _
+ _ _ _ _)))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ (_ @ _))))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Ident _ _ _ t9 idc9 @ @expr.LetIn _
+ _ _ _ _ _ _)))%expr_pat => None
+ | (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ #(_)))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
($_)%expr))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
@expr.Abs _ _ _ _ _ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
- (_ @ _)))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Ident _ _ _ t7 idc7) x12 @
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (($_)%expr @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.Abs _ _ _ _ _ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (_ @ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
+ (@expr.LetIn _ _ _ _ _ _ _ @ _)))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Ident _ _ _ t8 idc8) x13 @
@expr.LetIn _ _ _ _ _ _ _))%expr_pat =>
None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ ($_)%expr _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (_ @ _) _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.App _ _ _
- s11 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
+ | (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ ($_)%expr _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.Abs _ _ _ _ _ _) _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (_ @ _) _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.App _ _ _
+ s12 _ (@expr.LetIn _ _ _ _ _ _ _) _ @ _))%expr_pat =>
None
- | (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ #(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _
+ | (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ #(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ ($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Abs _ _ _
_ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
- (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _
+ (@expr.Ident _ _ _ t6 idc6 @
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.LetIn _ _
_ _ _ _ _ @ _))%expr_pat => None
- | (@expr.Ident _ _ _ t5 idc5 @ #(_))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ ($_)%expr)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.Abs _ _ _
+ | (@expr.Ident _ _ _ t6 idc6 @ #(_))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.Abs _ _ _
_ _ _)%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (#(_) @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (($_)%expr @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (#(_) @ _))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @ (($_)%expr @ _))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
(@expr.Abs _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
(($_)%expr @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
(@expr.Abs _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ (_ @ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @ (_ @ _ @ _ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t6 idc6 @
(@expr.LetIn _ _ _ _ _ _ _ @ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @
+ (@expr.Ident _ _ _ t6 idc6 @
(@expr.LetIn _ _ _ _ _ _ _ @ _))%expr_pat |
- (@expr.Ident _ _ _ t5 idc5 @ @expr.LetIn _ _
+ (@expr.Ident _ _ _ t6 idc6 @ @expr.LetIn _ _
_ _ _ _ _)%expr_pat => None
| _ => None
end;;
- match x8 with
- | @expr.App _ _ _ s8 _
- (@expr.Ident _ _ _ t5 idc5) x9 =>
- match x7 with
- | @expr.Ident _ _ _ t6 idc6 =>
- args <- invert_bind_args idc6
+ match x9 with
+ | @expr.App _ _ _ s9 _
+ (@expr.Ident _ _ _ t6 idc6) x10 =>
+ match x8 with
+ | (@expr.Ident _ _ _ t7 idc7 @ @expr.Ident
+ _ _ _ t8 idc8)%expr_pat =>
+ args <- invert_bind_args idc8
Raw.ident.Literal;
- args0 <- invert_bind_args idc5
+ args0 <- invert_bind_args idc7
+ Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc6
Raw.ident.Z_cast;
- _ <- invert_bind_args idc4
+ _ <- invert_bind_args idc5
Raw.ident.Z_shiftr;
- args2 <- invert_bind_args idc3
+ args3 <- invert_bind_args idc4
Raw.ident.Z_cast;
- args3 <- invert_bind_args idc2
+ args4 <- invert_bind_args idc3
Raw.ident.Z_cast;
- args4 <- invert_bind_args idc1
+ args5 <- invert_bind_args idc2
Raw.ident.Z_cast;
- args5 <- invert_bind_args idc0
+ args6 <- invert_bind_args idc1
Raw.ident.Literal;
+ args7 <- invert_bind_args idc0
+ Raw.ident.Z_cast;
_ <- invert_bind_args idc
Raw.ident.Z_sub_with_get_borrow;
match
pattern.type.unify_extracted
(((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) -> s3) -> s4) ->
- s8 -> (projT1 args))%ptype
+ ((((projT1 args6) -> s4) -> s5) ->
+ s9 -> (projT1 args))%ptype
with
| Some (_, _, _, (_, _))%zrange =>
if
type.type_beq base.type
base.type.type_beq
(((ℤ -> ℤ) -> ℤ) -> ℤ -> ℤ)%ptype
- ((((projT1 args5) -> s3) -> s4) ->
- s8 -> (projT1 args))%ptype
+ ((((projT1 args6) -> s4) -> s5) ->
+ s9 -> (projT1 args))%ptype
then
xv <- ident.unify
pattern.ident.Literal
- ##(projT2 args5);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s4 ℤ;
- v1 <- type.try_make_transport_cps s8 ℤ;
+ ##(projT2 args6);
+ v <- type.try_make_transport_cps s4 ℤ;
+ v0 <- type.try_make_transport_cps s5 ℤ;
+ v1 <- type.try_make_transport_cps s9 ℤ;
xv0 <- ident.unify
pattern.ident.Literal
##(projT2 args);
- fv <- (x10 <- (let
+ fv <- (x12 <- (let
'(r1, r2)%zrange :=
range in
- fun (s9 : Z)
+ fun (rs : zrange)
+ (s11 : Z)
(rb : zrange)
(b3 : expr ℤ)
(rx : zrange)
- (x10 : expr ℤ)
+ (x12 : expr ℤ)
(rshiftr
ry : zrange)
(y : expr ℤ)
+ (roffset : zrange)
(offset : Z) =>
if
- (s9 =?
- 2 ^ Z.log2 s9) &&
+ (s11 =?
+ 2 ^ Z.log2 s11) &&
(ZRange.normalize
ry >>
ZRange.normalize
@@ -9703,124 +6050,146 @@ match idc in (Compilers.ident t) return (Compile.value' true t) with
rshiftr)%zrange &&
(ZRange.normalize
rshiftr <=?
- r[0 ~> s9 - 1])%zrange
+ r[0 ~> s11 - 1])%zrange &&
+ is_bounded_by_bool
+ s11
+ (ZRange.normalize
+ rs) &&
+ is_bounded_by_bool
+ offset
+ (ZRange.normalize
+ roffset)
then
Some
(#(Z_cast2
(r1, r2))%expr @
(#(fancy_subb
(Z.log2
- s9)
+ s11)
(- offset))%expr @
(#(Z_cast rb)%expr @
b3,
#(Z_cast rx)%expr @
- x10,
+ x12,
#(Z_cast ry)%expr @
y)))%expr_pat
- else None)
- (let (x10, _) :=
+ else None) args7
+ (let (x12, _) :=
xv in
- x10) args4
+ x12) args5
(v
(Compile.reflect
- x4)) args3
+ x5)) args4
(v0
(Compile.reflect
- x5)) args2
- args0
+ x6)) args3
+ args1
(v1
(Compile.reflect
- x9))
- (let (x10, _) :=
+ x10)) args0
+ (let (x12, _) :=
xv0 in
- x10);
- Some (Base x10));
+ x12);
+ Some (Base x12));
Some
(fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
+ | (@expr.Ident _ _ _ t7 idc7 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.Abs _
+ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t7 idc7 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t7 idc7 @ @expr.LetIn
+ _ _ _ _ _ _ _)%expr_pat => None
| _ => None
end
- | @expr.App _ _ _ s8 _ ($_)%expr _ | @expr.App
- _ _ _ s8 _ (@expr.Abs _ _ _ _ _ _) _ |
- @expr.App _ _ _ s8 _ (_ @ _)%expr_pat _ |
- @expr.App _ _ _ s8 _
+ | @expr.App _ _ _ s9 _ ($_)%expr _ | @expr.App
+ _ _ _ s9 _ (@expr.Abs _ _ _ _ _ _) _ |
+ @expr.App _ _ _ s9 _ (_ @ _)%expr_pat _ |
+ @expr.App _ _ _ s9 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
| _ => None
end;;
- args <- invert_bind_args idc3 Raw.ident.Z_cast;
- args0 <- invert_bind_args idc2 Raw.ident.Z_cast;
- args1 <- invert_bind_args idc1 Raw.ident.Z_cast;
- args2 <- invert_bind_args idc0 Raw.ident.Literal;
+ args <- invert_bind_args idc4 Raw.ident.Z_cast;
+ args0 <- invert_bind_args idc3 Raw.ident.Z_cast;
+ args1 <- invert_bind_args idc2 Raw.ident.Z_cast;
+ args2 <- invert_bind_args idc1 Raw.ident.Literal;
+ args3 <- invert_bind_args idc0 Raw.ident.Z_cast;
_ <- invert_bind_args idc
Raw.ident.Z_sub_with_get_borrow;
match
pattern.type.unify_extracted
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> s4) -> s5)%ptype
+ ((((projT1 args2) -> s4) -> s5) -> s6)%ptype
with
| Some (_, _, _, _)%zrange =>
if
type.type_beq base.type base.type.type_beq
(((ℤ -> ℤ) -> ℤ) -> ℤ)%ptype
- ((((projT1 args2) -> s3) -> s4) -> s5)%ptype
+ ((((projT1 args2) -> s4) -> s5) -> s6)%ptype
then
xv <- ident.unify pattern.ident.Literal
##(projT2 args2);
- v <- type.try_make_transport_cps s3 ℤ;
- v0 <- type.try_make_transport_cps s4 ℤ;
- v1 <- type.try_make_transport_cps s5 ℤ;
- fv <- (x7 <- (let
+ v <- type.try_make_transport_cps s4 ℤ;
+ v0 <- type.try_make_transport_cps s5 ℤ;
+ v1 <- type.try_make_transport_cps s6 ℤ;
+ fv <- (x8 <- (let
'(r1, r2)%zrange := range in
- fun (s6 : Z) (rb : zrange)
- (b2 : expr ℤ) (rx : zrange)
- (x7 : expr ℤ) (ry : zrange)
- (y : expr ℤ) =>
+ fun (rs : zrange) (s7 : Z)
+ (rb : zrange) (b2 : expr ℤ)
+ (rx : zrange) (x8 : expr ℤ)
+ (ry : zrange) (y : expr ℤ) =>
if
- (s6 =? 2 ^ Z.log2 s6) &&
+ (s7 =? 2 ^ Z.log2 s7) &&
(ZRange.normalize ry <=?
- r[0 ~> s6 - 1])%zrange
+ r[0 ~> s7 - 1])%zrange &&
+ is_bounded_by_bool s7
+ (ZRange.normalize rs)
then
Some
(#(Z_cast2 (r1, r2))%expr @
- (#(fancy_subb (Z.log2 s6) 0)%expr @
+ (#(fancy_subb (Z.log2 s7) 0)%expr @
(#(Z_cast rb)%expr @ b2,
- #(Z_cast rx)%expr @ x7,
+ #(Z_cast rx)%expr @ x8,
#(Z_cast ry)%expr @ y)))%expr_pat
- else None)
- (let (x7, _) := xv in x7) args1
- (v (Compile.reflect x4)) args0
- (v0 (Compile.reflect x5)) args
- (v1 (Compile.reflect x6));
- Some (Base x7));
+ else None) args3
+ (let (x8, _) := xv in x8) args1
+ (v (Compile.reflect x5)) args0
+ (v0 (Compile.reflect x6)) args
+ (v1 (Compile.reflect x7));
+ Some (Base x8));
Some (fv0 <-- fv;
Base fv0)%under_lets
else None
| None => None
end
- | @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _ _ s5
- _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s5 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s5 _
+ | @expr.App _ _ _ s6 _ ($_)%expr _ | @expr.App _ _ _ s6
+ _ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s6 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s6 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
- | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s4 _
+ | @expr.App _ _ _ s5 _ ($_)%expr _ | @expr.App _ _ _ s5 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s5 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s5 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
- | @expr.App _ _ _ s3 _ ($_)%expr _ | @expr.App _ _ _ s3 _
- (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s3 _
- (_ @ _)%expr_pat _ | @expr.App _ _ _ s3 _
+ | @expr.App _ _ _ s4 _ ($_)%expr _ | @expr.App _ _ _ s4 _
+ (@expr.Abs _ _ _ _ _ _) _ | @expr.App _ _ _ s4 _
+ (_ @ _)%expr_pat _ | @expr.App _ _ _ s4 _
(@expr.LetIn _ _ _ _ _ _ _) _ => None
| _ => None
end
+ | (@expr.Ident _ _ _ t0 idc0 @ ($_)%expr)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.Abs _ _ _ _ _ _)%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ (_ @ _))%expr_pat |
+ (@expr.Ident _ _ _ t0 idc0 @ @expr.LetIn _ _ _ _ _ _ _)%expr_pat =>
+ None
| _ => None
end
| _ => None
diff --git a/src/strip_literal_casts_rewrite_head.out b/src/strip_literal_casts_rewrite_head.out
new file mode 100644
index 000000000..30c18d604
--- /dev/null
+++ b/src/strip_literal_casts_rewrite_head.out
@@ -0,0 +1,227 @@
+strip_literal_casts_rewrite_head =
+match idc in (Compilers.ident t) return (Compile.value' true t) with
+| @Literal t v => Base (##v)%expr
+| Nat_succ => fun x : expr ℕ => Base (#(Nat_succ)%expr @ x)%expr_pat
+| Nat_pred => fun x : expr ℕ => Base (#(Nat_pred)%expr @ x)%expr_pat
+| Nat_max => fun x x0 : expr ℕ => Base (#(Nat_max)%expr @ x @ x0)%expr_pat
+| Nat_mul => fun x x0 : expr ℕ => Base (#(Nat_mul)%expr @ x @ x0)%expr_pat
+| Nat_add => fun x x0 : expr ℕ => Base (#(Nat_add)%expr @ x @ x0)%expr_pat
+| Nat_sub => fun x x0 : expr ℕ => Base (#(Nat_sub)%expr @ x @ x0)%expr_pat
+| Nat_eqb => fun x x0 : expr ℕ => Base (#(Nat_eqb)%expr @ x @ x0)%expr_pat
+| @nil t => Base []%expr_pat
+| @cons t => fun (x : expr t) (x0 : expr (list t)) => Base (x :: x0)%expr_pat
+| @pair A B => fun (x : expr A) (x0 : expr B) => Base (x, x0)%expr_pat
+| @fst A B => fun x : expr (A * B)%etype => Base (#(fst)%expr @ x)%expr_pat
+| @snd A B => fun x : expr (A * B)%etype => Base (#(snd)%expr @ x)%expr_pat
+| @prod_rect A B T =>
+ fun (x : expr A -> expr B -> UnderLets (expr T))
+ (x0 : expr (A * B)%etype) =>
+ Base
+ (#(prod_rect)%expr @
+ (λ (x1 : var A)(x2 : var B),
+ to_expr (x ($x1) ($x2)))%expr @ x0)%expr_pat
+| @bool_rect T =>
+ fun (x x0 : expr unit -> UnderLets (expr T)) (x1 : expr bool) =>
+ Base
+ (#(bool_rect)%expr @ (λ x2 : var unit,
+ to_expr (x ($x2)))%expr @
+ (λ x2 : var unit,
+ to_expr (x0 ($x2)))%expr @ x1)%expr_pat
+| @nat_rect P =>
+ fun (x : expr unit -> UnderLets (expr P))
+ (x0 : expr ℕ -> expr P -> UnderLets (expr P)) (x1 : expr ℕ) =>
+ 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
+| @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) =>
+ Base
+ (#(nat_rect_arrow)%expr @ (λ x3 : var P,
+ to_expr (x ($x3)))%expr @
+ (λ (x3 : var ℕ)(x4 : var (P -> Q)%ptype)(x5 : var P),
+ to_expr
+ (x0 ($x3) (fun x6 : expr P => Base ($x4 @ x6)%expr_pat) ($x5)))%expr @
+ x1 @ x2)%expr_pat
+| @list_rect A P =>
+ fun (x : expr unit -> UnderLets (expr P))
+ (x0 : expr A -> expr (list A) -> expr P -> UnderLets (expr P))
+ (x1 : expr (list A)) =>
+ Base
+ (#(list_rect)%expr @ (λ x2 : var unit,
+ to_expr (x ($x2)))%expr @
+ (λ (x2 : var A)(x3 : var (list A))(x4 : var P),
+ to_expr (x0 ($x2) ($x3) ($x4)))%expr @ x1)%expr_pat
+| @list_case A P =>
+ fun (x : expr unit -> UnderLets (expr P))
+ (x0 : expr A -> expr (list A) -> UnderLets (expr P))
+ (x1 : expr (list A)) =>
+ Base
+ (#(list_case)%expr @ (λ x2 : var unit,
+ to_expr (x ($x2)))%expr @
+ (λ (x2 : var A)(x3 : var (list A)),
+ to_expr (x0 ($x2) ($x3)))%expr @ x1)%expr_pat
+| @List_length T =>
+ fun x : expr (list T) => Base (#(List_length)%expr @ x)%expr_pat
+| List_seq => fun x x0 : expr ℕ => Base (#(List_seq)%expr @ x @ x0)%expr_pat
+| @List_firstn A =>
+ fun (x : expr ℕ) (x0 : expr (list A)) =>
+ Base (#(List_firstn)%expr @ x @ x0)%expr_pat
+| @List_skipn A =>
+ fun (x : expr ℕ) (x0 : expr (list A)) =>
+ Base (#(List_skipn)%expr @ x @ x0)%expr_pat
+| @List_repeat A =>
+ fun (x : expr A) (x0 : expr ℕ) =>
+ Base (#(List_repeat)%expr @ x @ x0)%expr_pat
+| @List_combine A B =>
+ fun (x : expr (list A)) (x0 : expr (list B)) =>
+ Base (#(List_combine)%expr @ x @ x0)%expr_pat
+| @List_map A B =>
+ fun (x : expr A -> UnderLets (expr B)) (x0 : expr (list A)) =>
+ Base
+ (#(List_map)%expr @ (λ x1 : var A,
+ to_expr (x ($x1)))%expr @ x0)%expr_pat
+| @List_app A => fun x x0 : expr (list A) => Base (x ++ x0)%expr
+| @List_rev A =>
+ fun x : expr (list A) => Base (#(List_rev)%expr @ x)%expr_pat
+| @List_flat_map A B =>
+ fun (x : expr A -> UnderLets (expr (list B))) (x0 : expr (list A)) =>
+ Base
+ (#(List_flat_map)%expr @ (λ x1 : var A,
+ to_expr (x ($x1)))%expr @ x0)%expr_pat
+| @List_partition A =>
+ fun (x : expr A -> UnderLets (expr bool)) (x0 : expr (list A)) =>
+ Base
+ (#(List_partition)%expr @ (λ x1 : var A,
+ to_expr (x ($x1)))%expr @ x0)%expr_pat
+| @List_fold_right A B =>
+ fun (x : expr B -> expr A -> UnderLets (expr A)) (x0 : expr A)
+ (x1 : expr (list B)) =>
+ Base
+ (#(List_fold_right)%expr @
+ (λ (x2 : var B)(x3 : var A),
+ to_expr (x ($x2) ($x3)))%expr @ x0 @ x1)%expr_pat
+| @List_update_nth T =>
+ fun (x : expr ℕ) (x0 : expr T -> UnderLets (expr T)) (x1 : expr (list T))
+ =>
+ Base
+ (#(List_update_nth)%expr @ x @ (λ x2 : var T,
+ to_expr (x0 ($x2)))%expr @ x1)%expr_pat
+| @List_nth_default T =>
+ fun (x : expr T) (x0 : expr (list T)) (x1 : expr ℕ) =>
+ Base (#(List_nth_default)%expr @ x @ x0 @ x1)%expr_pat
+| Z_add => fun x x0 : expr ℤ => Base (x + x0)%expr
+| Z_mul => fun x x0 : expr ℤ => Base (x * x0)%expr
+| Z_pow => fun x x0 : expr ℤ => Base (#(Z_pow)%expr @ x @ x0)%expr_pat
+| Z_sub => fun x x0 : expr ℤ => Base (x - x0)%expr
+| Z_opp => fun x : expr ℤ => Base (- x)%expr
+| Z_div => fun x x0 : expr ℤ => Base (x / x0)%expr
+| Z_modulo => fun x x0 : expr ℤ => Base (x mod x0)%expr
+| Z_log2 => fun x : expr ℤ => Base (#(Z_log2)%expr @ x)%expr_pat
+| Z_log2_up => fun x : expr ℤ => Base (#(Z_log2_up)%expr @ x)%expr_pat
+| Z_eqb => fun x x0 : expr ℤ => Base (#(Z_eqb)%expr @ x @ x0)%expr_pat
+| Z_leb => fun x x0 : expr ℤ => Base (#(Z_leb)%expr @ x @ x0)%expr_pat
+| Z_geb => fun x x0 : expr ℤ => Base (#(Z_geb)%expr @ x @ x0)%expr_pat
+| Z_of_nat => fun x : expr ℕ => Base (#(Z_of_nat)%expr @ x)%expr_pat
+| Z_to_nat => fun x : expr ℤ => Base (#(Z_to_nat)%expr @ x)%expr_pat
+| Z_shiftr => fun x x0 : expr ℤ => Base (x >> x0)%expr
+| Z_shiftl => fun x x0 : expr ℤ => Base (x << x0)%expr
+| Z_land => fun x x0 : expr ℤ => Base (x &' x0)%expr
+| Z_lor => fun x x0 : expr ℤ => Base (x || x0)%expr
+| Z_bneg => fun x : expr ℤ => Base (#(Z_bneg)%expr @ x)%expr_pat
+| Z_lnot_modulo =>
+ fun x x0 : expr ℤ => Base (#(Z_lnot_modulo)%expr @ x @ x0)%expr_pat
+| Z_mul_split =>
+ fun x x0 x1 : expr ℤ => Base (#(Z_mul_split)%expr @ x @ x0 @ x1)%expr_pat
+| Z_add_get_carry =>
+ fun x x0 x1 : expr ℤ =>
+ Base (#(Z_add_get_carry)%expr @ x @ x0 @ x1)%expr_pat
+| Z_add_with_carry =>
+ fun x x0 x1 : expr ℤ =>
+ Base (#(Z_add_with_carry)%expr @ x @ x0 @ x1)%expr_pat
+| Z_add_with_get_carry =>
+ fun x x0 x1 x2 : expr ℤ =>
+ Base (#(Z_add_with_get_carry)%expr @ x @ x0 @ x1 @ x2)%expr_pat
+| Z_sub_get_borrow =>
+ fun x x0 x1 : expr ℤ =>
+ Base (#(Z_sub_get_borrow)%expr @ x @ x0 @ x1)%expr_pat
+| Z_sub_with_get_borrow =>
+ fun x x0 x1 x2 : expr ℤ =>
+ Base (#(Z_sub_with_get_borrow)%expr @ x @ x0 @ x1 @ x2)%expr_pat
+| Z_zselect =>
+ fun x x0 x1 : expr ℤ => Base (#(Z_zselect)%expr @ x @ x0 @ x1)%expr_pat
+| Z_add_modulo =>
+ fun x x0 x1 : expr ℤ =>
+ Base (#(Z_add_modulo)%expr @ x @ x0 @ x1)%expr_pat
+| Z_rshi =>
+ fun x x0 x1 x2 : expr ℤ =>
+ Base (#(Z_rshi)%expr @ x @ x0 @ x1 @ x2)%expr_pat
+| Z_cc_m => fun x x0 : expr ℤ => Base (#(Z_cc_m)%expr @ x @ x0)%expr_pat
+| Z_cast range =>
+ fun x : expr ℤ =>
+ (match x with
+ | @expr.Ident _ _ _ t idc =>
+ args <- invert_bind_args idc Raw.ident.Literal;
+ match pattern.type.unify_extracted ℤ (projT1 args) with
+ | Some _ =>
+ if type.type_beq base.type base.type.type_beq ℤ (projT1 args)
+ then
+ xv <- ident.unify pattern.ident.Literal ##(projT2 args);
+ fv <- (x0 <- (if
+ is_bounded_by_bool (let (x0, _) := xv in x0)
+ (ZRange.normalize range)
+ then Some (##(let (x0, _) := xv in x0))%expr
+ else None);
+ Some (Base x0));
+ Some (fv0 <-- fv;
+ Base fv0)%under_lets
+ else None
+ | None => None
+ end
+ | _ => None
+ end;;;
+ Base (#(Z_cast range)%expr @ x)%expr_pat)%option
+| Z_cast2 range =>
+ fun x : expr (ℤ * ℤ)%etype => Base (#(Z_cast2 range)%expr @ x)%expr_pat
+| fancy_add log2wordmax imm =>
+ fun x : expr (ℤ * ℤ)%etype =>
+ Base (#(fancy_add log2wordmax imm)%expr @ x)%expr_pat
+| fancy_addc log2wordmax imm =>
+ fun x : expr (ℤ * ℤ * ℤ)%etype =>
+ Base (#(fancy_addc log2wordmax imm)%expr @ x)%expr_pat
+| fancy_sub log2wordmax imm =>
+ fun x : expr (ℤ * ℤ)%etype =>
+ Base (#(fancy_sub log2wordmax imm)%expr @ x)%expr_pat
+| fancy_subb log2wordmax imm =>
+ fun x : expr (ℤ * ℤ * ℤ)%etype =>
+ Base (#(fancy_subb log2wordmax imm)%expr @ x)%expr_pat
+| fancy_mulll log2wordmax =>
+ fun x : expr (ℤ * ℤ)%etype =>
+ Base (#(fancy_mulll log2wordmax)%expr @ x)%expr_pat
+| fancy_mullh log2wordmax =>
+ fun x : expr (ℤ * ℤ)%etype =>
+ Base (#(fancy_mullh log2wordmax)%expr @ x)%expr_pat
+| fancy_mulhl log2wordmax =>
+ fun x : expr (ℤ * ℤ)%etype =>
+ Base (#(fancy_mulhl log2wordmax)%expr @ x)%expr_pat
+| fancy_mulhh log2wordmax =>
+ fun x : expr (ℤ * ℤ)%etype =>
+ Base (#(fancy_mulhh log2wordmax)%expr @ x)%expr_pat
+| fancy_rshi log2wordmax x =>
+ fun x0 : expr (ℤ * ℤ)%etype =>
+ Base (#(fancy_rshi log2wordmax x)%expr @ x0)%expr_pat
+| fancy_selc =>
+ fun x : expr (ℤ * ℤ * ℤ)%etype => Base (#(fancy_selc)%expr @ x)%expr_pat
+| fancy_selm log2wordmax =>
+ fun x : expr (ℤ * ℤ * ℤ)%etype =>
+ Base (#(fancy_selm log2wordmax)%expr @ x)%expr_pat
+| fancy_sell =>
+ fun x : expr (ℤ * ℤ * ℤ)%etype => Base (#(fancy_sell)%expr @ x)%expr_pat
+| fancy_addm =>
+ fun x : expr (ℤ * ℤ * ℤ)%etype => Base (#(fancy_addm)%expr @ x)%expr_pat
+end
+ : Compile.value' true t