aboutsummaryrefslogtreecommitdiff
path: root/src/Compilers/Z
diff options
context:
space:
mode:
authorGravatar Andres Erbsen <andreser@mit.edu>2017-04-06 22:53:07 -0400
committerGravatar Andres Erbsen <andreser@mit.edu>2017-04-06 22:53:07 -0400
commitc9fc5a3cdf1f5ea2d104c150c30d1b1a6ac64239 (patch)
treedb7187f6984acff324ca468e7b33d9285806a1eb /src/Compilers/Z
parent21198245dab432d3c0ba2bb8a02254e7d0594382 (diff)
rename-everything
Diffstat (limited to 'src/Compilers/Z')
-rw-r--r--src/Compilers/Z/ArithmeticSimplifier.v184
-rw-r--r--src/Compilers/Z/ArithmeticSimplifierInterp.v120
-rw-r--r--src/Compilers/Z/ArithmeticSimplifierUtil.v79
-rw-r--r--src/Compilers/Z/ArithmeticSimplifierWf.v168
-rw-r--r--src/Compilers/Z/BinaryNotationConstants.v91
-rw-r--r--src/Compilers/Z/Bounds/Interpretation.v177
-rw-r--r--src/Compilers/Z/Bounds/InterpretationLemmas.v433
-rw-r--r--src/Compilers/Z/Bounds/MapCastByDeBruijn.v23
-rw-r--r--src/Compilers/Z/Bounds/MapCastByDeBruijnInterp.v25
-rw-r--r--src/Compilers/Z/Bounds/MapCastByDeBruijnWf.v41
-rw-r--r--src/Compilers/Z/Bounds/Pipeline.v20
-rw-r--r--src/Compilers/Z/Bounds/Pipeline/Definition.v177
-rw-r--r--src/Compilers/Z/Bounds/Pipeline/Glue.v456
-rw-r--r--src/Compilers/Z/Bounds/Pipeline/OutputType.v51
-rw-r--r--src/Compilers/Z/Bounds/Pipeline/ReflectiveTactics.v288
-rw-r--r--src/Compilers/Z/Bounds/Relax.v127
-rw-r--r--src/Compilers/Z/CNotations.v773
-rw-r--r--src/Compilers/Z/FoldTypes.v17
-rw-r--r--src/Compilers/Z/HexNotationConstants.v144
-rw-r--r--src/Compilers/Z/Inline.v7
-rw-r--r--src/Compilers/Z/InlineInterp.v11
-rw-r--r--src/Compilers/Z/InlineWf.v11
-rw-r--r--src/Compilers/Z/JavaNotations.v792
-rw-r--r--src/Compilers/Z/MapCastByDeBruijn.v28
-rw-r--r--src/Compilers/Z/MapCastByDeBruijnInterp.v50
-rw-r--r--src/Compilers/Z/MapCastByDeBruijnWf.v56
-rw-r--r--src/Compilers/Z/OpInversion.v28
-rw-r--r--src/Compilers/Z/Reify.v50
-rw-r--r--src/Compilers/Z/Syntax.v84
-rw-r--r--src/Compilers/Z/Syntax/Equality.v176
-rw-r--r--src/Compilers/Z/Syntax/Util.v170
31 files changed, 4857 insertions, 0 deletions
diff --git a/src/Compilers/Z/ArithmeticSimplifier.v b/src/Compilers/Z/ArithmeticSimplifier.v
new file mode 100644
index 000000000..b2621c625
--- /dev/null
+++ b/src/Compilers/Z/ArithmeticSimplifier.v
@@ -0,0 +1,184 @@
+(** * SimplifyArith: Remove things like (_ * 1), (_ + 0), etc *)
+Require Import Coq.ZArith.ZArith.
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Rewriter.
+Require Import Crypto.Compilers.Z.Syntax.
+
+Section language.
+ Local Notation exprf := (@exprf base_type op).
+ Local Notation Expr := (@Expr base_type op).
+
+ Section with_var.
+ Context {var : base_type -> Type}.
+
+ Inductive inverted_expr t :=
+ | const_of (v : Z)
+ | gen_expr (e : exprf (var:=var) (Tbase t))
+ | neg_expr (e : exprf (var:=var) (Tbase t)).
+
+ Fixpoint interp_as_expr_or_const {t} (x : exprf (var:=var) t)
+ : option (interp_flat_type inverted_expr t)
+ := match x in Syntax.exprf _ _ t return option (interp_flat_type _ t) with
+ | Op t1 (Tbase _) opc args
+ => Some (match opc in op src dst return exprf dst -> exprf src -> inverted_expr match dst with Tbase t => t | _ => TZ end with
+ | OpConst _ z => fun _ _ => const_of _ z
+ | Opp TZ TZ => fun _ args => neg_expr _ args
+ | _ => fun e _ => gen_expr _ e
+ end (Op opc args) args)
+ | TT => Some tt
+ | Var t v => Some (gen_expr _ (Var v))
+ | Op _ _ _ _
+ | LetIn _ _ _ _
+ => None
+ | Pair tx ex ty ey
+ => match @interp_as_expr_or_const tx ex, @interp_as_expr_or_const ty ey with
+ | Some vx, Some vy => Some (vx, vy)
+ | _, None | None, _ => None
+ end
+ end.
+
+ Definition simplify_op_expr {src dst} (opc : op src dst)
+ : exprf (var:=var) src -> exprf (var:=var) dst
+ := match opc in op src dst return exprf src -> exprf dst with
+ | Add TZ TZ TZ as opc
+ => fun args
+ => match interp_as_expr_or_const args with
+ | Some (const_of l, const_of r)
+ => Op (OpConst (interp_op _ _ opc (l, r))) TT
+ | Some (const_of v, gen_expr e)
+ | Some (gen_expr e, const_of v)
+ => if (v =? 0)%Z
+ then e
+ else Op opc args
+ | Some (const_of v, neg_expr e)
+ | Some (neg_expr e, const_of v)
+ => if (v =? 0)%Z
+ then Op (Opp _ _) e
+ else Op opc args
+ | Some (gen_expr ep, neg_expr en)
+ | Some (neg_expr en, gen_expr ep)
+ => Op (Sub _ _ _) (Pair ep en)
+ | _ => Op opc args
+ end
+ | Sub TZ TZ TZ as opc
+ => fun args
+ => match interp_as_expr_or_const args with
+ | Some (const_of l, const_of r)
+ => Op (OpConst (interp_op _ _ opc (l, r))) TT
+ | Some (gen_expr e, const_of v)
+ => if (v =? 0)%Z
+ then e
+ else Op opc args
+ | Some (neg_expr e, const_of v)
+ => if (v =? 0)%Z
+ then Op (Opp _ _) e
+ else Op opc args
+ | Some (gen_expr e1, neg_expr e2)
+ => Op (Add _ _ _) (Pair e1 e2)
+ | Some (neg_expr e1, neg_expr e2)
+ => Op (Sub _ _ _) (Pair e2 e1)
+ | _ => Op opc args
+ end
+ | Mul TZ TZ TZ as opc
+ => fun args
+ => match interp_as_expr_or_const args with
+ | Some (const_of l, const_of r)
+ => Op (OpConst (interp_op _ _ opc (l, r))) TT
+ | Some (const_of v, gen_expr e)
+ | Some (gen_expr e, const_of v)
+ => if (v =? 0)%Z
+ then Op (OpConst 0%Z) TT
+ else if (v =? 1)%Z
+ then e
+ else if (v =? -1)%Z
+ then Op (Opp _ _) e
+ else Op opc args
+ | Some (const_of v, neg_expr e)
+ | Some (neg_expr e, const_of v)
+ => if (v =? 0)%Z
+ then Op (OpConst 0%Z) TT
+ else if (v =? 1)%Z
+ then Op (Opp _ _) e
+ else if (v =? -1)%Z
+ then e
+ else Op opc args
+ | Some (gen_expr e1, neg_expr e2)
+ | Some (neg_expr e1, gen_expr e2)
+ => Op (Opp _ _) (Op (Mul _ _ TZ) (Pair e1 e2))
+ | Some (neg_expr e1, neg_expr e2)
+ => Op (Mul _ _ _) (Pair e1 e2)
+ | _ => Op opc args
+ end
+ | Shl TZ TZ TZ as opc
+ | Shr TZ TZ TZ as opc
+ => fun args
+ => match interp_as_expr_or_const args with
+ | Some (const_of l, const_of r)
+ => Op (OpConst (interp_op _ _ opc (l, r))) TT
+ | Some (gen_expr e, const_of v)
+ => if (v =? 0)%Z
+ then e
+ else Op opc args
+ | Some (neg_expr e, const_of v)
+ => if (v =? 0)%Z
+ then Op (Opp _ _) e
+ else Op opc args
+ | _ => Op opc args
+ end
+ | Land TZ TZ TZ as opc
+ => fun args
+ => match interp_as_expr_or_const args with
+ | Some (const_of l, const_of r)
+ => Op (OpConst (interp_op _ _ opc (l, r))) TT
+ | Some (const_of v, gen_expr _)
+ | Some (gen_expr _, const_of v)
+ | Some (const_of v, neg_expr _)
+ | Some (neg_expr _, const_of v)
+ => if (v =? 0)%Z
+ then Op (OpConst 0%Z) TT
+ else Op opc args
+ | _ => Op opc args
+ end
+ | Lor TZ TZ TZ as opc
+ => fun args
+ => match interp_as_expr_or_const args with
+ | Some (const_of l, const_of r)
+ => Op (OpConst (interp_op _ _ opc (l, r))) TT
+ | Some (const_of v, gen_expr e)
+ | Some (gen_expr e, const_of v)
+ => if (v =? 0)%Z
+ then e
+ else Op opc args
+ | Some (const_of v, neg_expr e)
+ | Some (neg_expr e, const_of v)
+ => if (v =? 0)%Z
+ then Op (Opp _ _) e
+ else Op opc args
+ | _ => Op opc args
+ end
+ | Opp TZ TZ as opc
+ => fun args
+ => match interp_as_expr_or_const args with
+ | Some (const_of v)
+ => Op (OpConst (interp_op _ _ opc v)) TT
+ | Some (neg_expr e)
+ => e
+ | _
+ => Op opc args
+ end
+ | Add _ _ _ as opc
+ | Sub _ _ _ as opc
+ | Mul _ _ _ as opc
+ | Shl _ _ _ as opc
+ | Shr _ _ _ as opc
+ | Land _ _ _ as opc
+ | Lor _ _ _ as opc
+ | OpConst _ _ as opc
+ | Opp _ _ as opc
+ => Op opc
+ end.
+ End with_var.
+
+ Definition SimplifyArith {t} (e : Expr t) : Expr t
+ := @RewriteOp base_type op (@simplify_op_expr) t e.
+End language.
diff --git a/src/Compilers/Z/ArithmeticSimplifierInterp.v b/src/Compilers/Z/ArithmeticSimplifierInterp.v
new file mode 100644
index 000000000..6eec2f2a4
--- /dev/null
+++ b/src/Compilers/Z/ArithmeticSimplifierInterp.v
@@ -0,0 +1,120 @@
+Require Import Coq.micromega.Psatz.
+Require Import Coq.ZArith.ZArith.
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.TypeInversion.
+Require Import Crypto.Compilers.ExprInversion.
+Require Import Crypto.Compilers.RewriterInterp.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.OpInversion.
+Require Import Crypto.Compilers.Z.ArithmeticSimplifier.
+Require Import Crypto.Compilers.Z.ArithmeticSimplifierUtil.
+Require Import Crypto.Compilers.Z.Syntax.Equality.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Option.
+Require Import Crypto.Util.Prod.
+Require Import Crypto.Util.Sum.
+Require Import Crypto.Util.HProp.
+Require Import Crypto.Util.Tactics.BreakMatch.
+Require Import Crypto.Util.Tactics.DestructHead.
+
+Local Notation exprf := (@exprf base_type op interp_base_type).
+Local Notation expr := (@expr base_type op interp_base_type).
+Local Notation Expr := (@Expr base_type op).
+
+Local Ltac fin_t :=
+ first [ exact I
+ | reflexivity
+ | congruence
+ | assumption
+ | lia
+ | exfalso; assumption ].
+Local Ltac break_t_step :=
+ first [ progress subst
+ | progress inversion_option
+ | progress inversion_sum
+ | progress inversion_expr
+ | progress inversion_prod
+ | progress inversion_inverted_expr
+ | progress inversion_flat_type
+ | progress destruct_head'_and
+ | progress destruct_head'_prod
+ | progress eliminate_hprop_eq
+ | progress break_innermost_match_step
+ | progress break_match_hyps ].
+
+
+Lemma interp_as_expr_or_const_correct_base {t} e z
+ : @interp_as_expr_or_const interp_base_type (Tbase t) e = Some z
+ -> interpf interp_op e = match z with
+ | const_of z => cast_const (t1:=TZ) z
+ | gen_expr e => interpf interp_op e
+ | neg_expr e => interpf interp_op (Op (Opp _ _) e)
+ end.
+Proof.
+ destruct z.
+ { repeat first [ fin_t
+ | progress simpl in *
+ | progress intros
+ | break_t_step
+ | progress invert_expr
+ | progress invert_op ]. }
+ { do 2 (invert_expr; break_innermost_match; intros);
+ repeat first [ fin_t
+ | progress simpl in *
+ | progress intros
+ | break_t_step
+ | progress invert_op ]. }
+ { do 2 (invert_expr; break_innermost_match; intros);
+ repeat first [ fin_t
+ | progress simpl in *
+ | progress intros
+ | break_t_step
+ | progress invert_op ]. }
+Qed.
+
+Lemma interp_as_expr_or_const_correct_prod_base {A B} e (v : _ * _)
+ : @interp_as_expr_or_const interp_base_type (Prod (Tbase A) (Tbase B)) e = Some v
+ -> interpf interp_op e = (match fst v with
+ | const_of z => cast_const (t1:=TZ) z
+ | gen_expr e => interpf interp_op e
+ | neg_expr e => interpf interp_op (Op (Opp _ _) e)
+ end,
+ match snd v with
+ | const_of z => cast_const (t1:=TZ) z
+ | gen_expr e => interpf interp_op e
+ | neg_expr e => interpf interp_op (Op (Opp _ _) e)
+ end).
+Proof.
+ invert_expr;
+ repeat first [ fin_t
+ | progress simpl in *
+ | progress intros
+ | break_t_step
+ | erewrite !interp_as_expr_or_const_correct_base by eassumption; cbv beta iota ].
+Qed.
+
+Local Arguments Z.mul !_ !_.
+Local Arguments Z.add !_ !_.
+Local Arguments Z.sub !_ !_.
+Local Arguments Z.opp !_.
+
+Lemma InterpSimplifyArith {t} (e : Expr t)
+ : forall x, Interp interp_op (SimplifyArith e) x = Interp interp_op e x.
+Proof.
+ apply InterpRewriteOp; intros; unfold simplify_op_expr.
+ break_innermost_match;
+ repeat first [ fin_t
+ | progress simpl in *
+ | progress subst
+ | erewrite !interp_as_expr_or_const_correct_prod_base by eassumption; cbv beta iota
+ | erewrite !interp_as_expr_or_const_correct_base by eassumption; cbv beta iota
+ | match goal with
+ | [ |- context[interpf _ ?e] ]
+ => erewrite !(@interp_as_expr_or_const_correct_base _ e) by eassumption; cbv beta iota
+ end
+ | progress unfold interp_op, lift_op
+ | progress Z.ltb_to_lt
+ | progress rewrite ?Z.land_0_l, ?Z.land_0_r, ?Z.lor_0_l, ?Z.lor_0_r ].
+Qed.
+
+Hint Rewrite @InterpSimplifyArith : reflective_interp.
diff --git a/src/Compilers/Z/ArithmeticSimplifierUtil.v b/src/Compilers/Z/ArithmeticSimplifierUtil.v
new file mode 100644
index 000000000..49d3a2257
--- /dev/null
+++ b/src/Compilers/Z/ArithmeticSimplifierUtil.v
@@ -0,0 +1,79 @@
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.ArithmeticSimplifier.
+
+(** ** Equality for [inverted_expr] *)
+Section inverted_expr.
+ Context {var : base_type -> Type}.
+ Local Notation inverted_expr := (@inverted_expr var).
+ Local Notation inverted_expr_code u v
+ := (match u, v with
+ | const_of u', const_of v'
+ | gen_expr u', gen_expr v'
+ | neg_expr u', neg_expr v'
+ => u' = v'
+ | const_of _, _
+ | gen_expr _, _
+ | neg_expr _, _
+ => False
+ end).
+
+ (** *** Equality of [inverted_expr] is a [match] *)
+ Definition path_inverted_expr {T} (u v : inverted_expr T) (p : inverted_expr_code u v)
+ : u = v.
+ Proof. destruct u, v; first [ apply f_equal | exfalso ]; exact p. Defined.
+
+ (** *** Equivalence of equality of [inverted_expr] with [inverted_expr_code] *)
+ Definition unpath_inverted_expr {T} {u v : inverted_expr T} (p : u = v)
+ : inverted_expr_code u v.
+ Proof. subst v; destruct u; reflexivity. Defined.
+
+ Definition path_inverted_expr_iff {T}
+ (u v : @inverted_expr T)
+ : u = v <-> inverted_expr_code u v.
+ Proof.
+ split; [ apply unpath_inverted_expr | apply path_inverted_expr ].
+ Defined.
+
+ (** *** Eta-expansion of [@eq (inverted_expr _ _)] *)
+ Definition path_inverted_expr_eta {T} {u v : @inverted_expr T} (p : u = v)
+ : p = path_inverted_expr u v (unpath_inverted_expr p).
+ Proof. destruct u, p; reflexivity. Defined.
+
+ (** *** Induction principle for [@eq (inverted_expr _ _)] *)
+ Definition path_inverted_expr_rect {T} {u v : @inverted_expr T} (P : u = v -> Type)
+ (f : forall p, P (path_inverted_expr u v p))
+ : forall p, P p.
+ Proof. intro p; specialize (f (unpath_inverted_expr p)); destruct u, p; exact f. Defined.
+ Definition path_inverted_expr_rec {T u v} (P : u = v :> @inverted_expr T -> Set) := path_inverted_expr_rect P.
+ Definition path_inverted_expr_ind {T u v} (P : u = v :> @inverted_expr T -> Prop) := path_inverted_expr_rec P.
+End inverted_expr.
+
+(** ** Useful Tactics *)
+(** *** [inversion_inverted_expr] *)
+Ltac induction_path_inverted_expr H :=
+ induction H as [H] using path_inverted_expr_rect;
+ try match type of H with
+ | False => exfalso; exact H
+ end.
+Ltac inversion_inverted_expr_step :=
+ match goal with
+ | [ H : const_of _ _ = const_of _ _ |- _ ]
+ => induction_path_inverted_expr H
+ | [ H : const_of _ _ = gen_expr _ _ |- _ ]
+ => induction_path_inverted_expr H
+ | [ H : const_of _ _ = neg_expr _ _ |- _ ]
+ => induction_path_inverted_expr H
+ | [ H : gen_expr _ _ = const_of _ _ |- _ ]
+ => induction_path_inverted_expr H
+ | [ H : gen_expr _ _ = gen_expr _ _ |- _ ]
+ => induction_path_inverted_expr H
+ | [ H : gen_expr _ _ = neg_expr _ _ |- _ ]
+ => induction_path_inverted_expr H
+ | [ H : neg_expr _ _ = const_of _ _ |- _ ]
+ => induction_path_inverted_expr H
+ | [ H : neg_expr _ _ = gen_expr _ _ |- _ ]
+ => induction_path_inverted_expr H
+ | [ H : neg_expr _ _ = neg_expr _ _ |- _ ]
+ => induction_path_inverted_expr H
+ end.
+Ltac inversion_inverted_expr := repeat inversion_inverted_expr_step.
diff --git a/src/Compilers/Z/ArithmeticSimplifierWf.v b/src/Compilers/Z/ArithmeticSimplifierWf.v
new file mode 100644
index 000000000..24a9b4d23
--- /dev/null
+++ b/src/Compilers/Z/ArithmeticSimplifierWf.v
@@ -0,0 +1,168 @@
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Wf.
+Require Import Crypto.Compilers.WfInversion.
+Require Import Crypto.Compilers.TypeInversion.
+Require Import Crypto.Compilers.ExprInversion.
+Require Import Crypto.Compilers.RewriterWf.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.OpInversion.
+Require Import Crypto.Compilers.Z.ArithmeticSimplifier.
+Require Import Crypto.Compilers.Z.Syntax.Equality.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Option.
+Require Import Crypto.Util.Sum.
+Require Import Crypto.Util.Prod.
+Require Import Crypto.Util.Tactics.BreakMatch.
+Require Import Crypto.Util.Tactics.DestructHead.
+Require Import Crypto.Util.Tactics.SpecializeBy.
+Require Import Crypto.Util.HProp.
+
+Local Notation exprf := (@exprf base_type op).
+Local Notation expr := (@expr base_type op).
+Local Notation Expr := (@Expr base_type op).
+Local Notation wff := (@wff base_type op).
+Local Notation Wf := (@Wf base_type op).
+
+Local Ltac fin_t :=
+ first [ exact I
+ | reflexivity
+ | congruence
+ | assumption
+ | exfalso; assumption ].
+Local Ltac break_t_step :=
+ first [ progress subst
+ | progress inversion_option
+ | progress inversion_sum
+ | progress inversion_expr
+ | progress inversion_prod
+ | progress invert_op
+ | progress inversion_flat_type
+ | progress destruct_head'_and
+ | progress destruct_head' iff
+ | progress destruct_head'_prod
+ | progress destruct_head'_sig
+ | progress specialize_by reflexivity
+ | progress eliminate_hprop_eq
+ | progress break_innermost_match_step
+ | progress break_match_hyps
+ | progress inversion_wf_constr ].
+
+
+Lemma interp_as_expr_or_const_None_iff {var1 var2 t} {G e1 e2}
+ (Hwf : @wff var1 var2 G t e1 e2)
+ : @interp_as_expr_or_const var1 t e1 = None
+ <-> @interp_as_expr_or_const var2 t e2 = None.
+Proof.
+ induction Hwf;
+ repeat first [ fin_t
+ | split; congruence
+ | progress simpl in *
+ | progress intros
+ | break_t_step ].
+Qed.
+
+Lemma interp_as_expr_or_const_None_Some {var1 var2 t} {G e1 e2 v}
+ (Hwf : @wff var1 var2 G t e1 e2)
+ : @interp_as_expr_or_const var1 t e1 = None
+ -> @interp_as_expr_or_const var2 t e2 = Some v
+ -> False.
+Proof.
+ erewrite interp_as_expr_or_const_None_iff by eassumption; congruence.
+Qed.
+
+Lemma interp_as_expr_or_const_Some_None {var1 var2 t} {G e1 e2 v}
+ (Hwf : @wff var1 var2 G t e1 e2)
+ : @interp_as_expr_or_const var1 t e1 = Some v
+ -> @interp_as_expr_or_const var2 t e2 = None
+ -> False.
+Proof.
+ erewrite <- interp_as_expr_or_const_None_iff by eassumption; congruence.
+Qed.
+
+Lemma wff_interp_as_expr_or_const_base {var1 var2 t} {G e1 e2 v1 v2}
+ (Hwf : @wff var1 var2 G (Tbase t) e1 e2)
+ : @interp_as_expr_or_const var1 (Tbase t) e1 = Some v1
+ -> @interp_as_expr_or_const var2 (Tbase t) e2 = Some v2
+ -> match v1, v2 with
+ | const_of z1, const_of z2 => z1 = z2
+ | gen_expr e1, gen_expr e2
+ | neg_expr e1, neg_expr e2
+ => wff G e1 e2
+ | const_of _, _
+ | gen_expr _, _
+ | neg_expr _, _
+ => False
+ end.
+Proof.
+ invert_one_expr e1; intros; break_innermost_match; intros;
+ try invert_one_expr e2; intros;
+ repeat first [ fin_t
+ | progress simpl in *
+ | progress intros
+ | break_t_step
+ | match goal with
+ | [ H : wff _ _ ?e |- _ ] => is_var e; invert_one_expr e
+ end ].
+Qed.
+
+Lemma wff_interp_as_expr_or_const_prod_base {var1 var2 A B} {G e1 e2} {v1 v2 : _ * _}
+ (Hwf : wff G e1 e2)
+ : @interp_as_expr_or_const var1 (Prod (Tbase A) (Tbase B)) e1 = Some v1
+ -> @interp_as_expr_or_const var2 (Prod (Tbase A) (Tbase B)) e2 = Some v2
+ -> match fst v1, fst v2 with
+ | const_of z1, const_of z2 => z1 = z2
+ | gen_expr e1, gen_expr e2
+ | neg_expr e1, neg_expr e2
+ => wff G e1 e2
+ | const_of _, _
+ | gen_expr _, _
+ | neg_expr _, _
+ => False
+ end
+ /\ match snd v1, snd v2 with
+ | const_of z1, const_of z2 => z1 = z2
+ | gen_expr e1, gen_expr e2
+ | neg_expr e1, neg_expr e2
+ => wff G e1 e2
+ | const_of _, _
+ | gen_expr _, _
+ | neg_expr _, _
+ => False
+ end.
+Proof.
+ invert_one_expr e1; intros; break_innermost_match; intros; try exact I;
+ try invert_one_expr e2; intros; break_innermost_match; intros; try exact I;
+ repeat first [ fin_t
+ | progress simpl in *
+ | break_t_step
+ | match goal with
+ | [ H1 : _ = Some _, H2 : _ = Some _, Hwf : wff _ _ _ |- _ ]
+ => pose proof (wff_interp_as_expr_or_const_base Hwf H1 H2); clear H1 H2
+ | [ |- and _ _ ] => split
+ end ].
+Qed.
+
+Lemma Wf_SimplifyArith {t} (e : Expr t)
+ (Hwf : Wf e)
+ : Wf (SimplifyArith e).
+Proof.
+ apply Wf_RewriteOp; [ | assumption ].
+ intros ???????? Hwf'; unfold simplify_op_expr;
+ break_innermost_match; repeat constructor; auto;
+ repeat first [ fin_t
+ | progress simpl in *
+ | progress subst
+ | progress Z.ltb_to_lt
+ | match goal with
+ | [ H1 : _ = Some _, H2 : _ = Some _, Hwf : wff _ _ _ |- _ ]
+ => pose proof (wff_interp_as_expr_or_const_base Hwf H1 H2); clear H1 H2
+ | [ H1 : _ = Some _, H2 : _ = Some _, Hwf : wff _ _ _ |- _ ]
+ => pose proof (wff_interp_as_expr_or_const_prod_base Hwf H1 H2); clear H1 H2
+ | [ H1 : _ = Some _, H2 : _ = None, Hwf : wff _ _ _ |- _ ]
+ => pose proof (interp_as_expr_or_const_Some_None Hwf H1 H2); clear H1 H2
+ | [ H1 : _ = None, H2 : _ = Some _, Hwf : wff _ _ _ |- _ ]
+ => pose proof (interp_as_expr_or_const_None_Some Hwf H1 H2); clear H1 H2
+ | [ |- wff _ _ _ ] => constructor
+ end
+ | break_t_step ].
+Qed.
diff --git a/src/Compilers/Z/BinaryNotationConstants.v b/src/Compilers/Z/BinaryNotationConstants.v
new file mode 100644
index 000000000..f26364329
--- /dev/null
+++ b/src/Compilers/Z/BinaryNotationConstants.v
@@ -0,0 +1,91 @@
+Require Export Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Export Bedrock.Word.
+Require Export Crypto.Util.Notations.
+
+Local Notation Const x := (Op (OpConst x) TT).
+(* python:
+<<
+print('\n'.join('''Notation "'%s'" (* %d (%s) *)\n := (Const %s).''' % (b, d, h, i)
+ for d, h, b, i in sorted([(eval(bv), hex(eval(bv)), bv, i)
+ for (bv, i) in (('0b' + i[2:].replace('~', ''), i)
+ for i in r"""WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1
+WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~1~0
+WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0
+WO~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1
+WO~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0
+WO~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1
+WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~1~0
+WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~0
+WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0
+WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1
+WO~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0
+WO~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1
+WO~0~0~0~1~0~0~1~1
+WO~0~0~0~1~1~0~0~1
+WO~0~0~0~1~1~0~1~0
+WO~0~0~0~1~1~0~1~1
+WO~0~0~0~1~1~1~0~0
+WO~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0
+WO~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0
+WO~0~0~1~1~0~0~1~1
+WO~1~0
+WO~1~0~0~1
+WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1
+WO~0~0~0~1~0~0~0~1
+WO~0~0~0~1~0~1~1~1
+WO~1~1""".split('\n'))])))
+>>
+ *)
+Notation "'0b10'" (* 2 (0x2) *)
+ := (Const WO~1~0).
+Notation "'0b11'" (* 3 (0x3) *)
+ := (Const WO~1~1).
+Notation "'0b1001'" (* 9 (0x9) *)
+ := (Const WO~1~0~0~1).
+Notation "'0b00010001'" (* 17 (0x11) *)
+ := (Const WO~0~0~0~1~0~0~0~1).
+Notation "'0b00010011'" (* 19 (0x13) *)
+ := (Const WO~0~0~0~1~0~0~1~1).
+Notation "'0b00010111'" (* 23 (0x17) *)
+ := (Const WO~0~0~0~1~0~1~1~1).
+Notation "'0b00011001'" (* 25 (0x19) *)
+ := (Const WO~0~0~0~1~1~0~0~1).
+Notation "'0b00011010'" (* 26 (0x1a) *)
+ := (Const WO~0~0~0~1~1~0~1~0).
+Notation "'0b00011011'" (* 27 (0x1b) *)
+ := (Const WO~0~0~0~1~1~0~1~1).
+Notation "'0b00011100'" (* 28 (0x1c) *)
+ := (Const WO~0~0~0~1~1~1~0~0).
+Notation "'0b00110011'" (* 51 (0x33) *)
+ := (Const WO~0~0~1~1~0~0~1~1).
+Notation "'0b00000000011111111111111111111111'" (* 8388607 (0x7fffff) *)
+ := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
+Notation "'0b00000001111111111111111111111111'" (* 33554431 (0x1ffffff) *)
+ := (Const WO~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
+Notation "'0b00000011111111111111111111111110'" (* 67108862 (0x3fffffe) *)
+ := (Const WO~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
+Notation "'0b00000011111111111111111111111111'" (* 67108863 (0x3ffffff) *)
+ := (Const WO~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
+Notation "'0b00000111111111111111111111011010'" (* 134217690 (0x7ffffda) *)
+ := (Const WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~1~0).
+Notation "'0b00000111111111111111111111101110'" (* 134217710 (0x7ffffee) *)
+ := (Const WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~0).
+Notation "'0b00000111111111111111111111111110'" (* 134217726 (0x7fffffe) *)
+ := (Const WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
+Notation "'0b00000111111111111111111111111111'" (* 134217727 (0x7ffffff) *)
+ := (Const WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
+Notation "'0b00001111111111111111111111111110'" (* 268435454 (0xffffffe) *)
+ := (Const WO~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
+Notation "'0b00001111111111111111111111111111'" (* 268435455 (0xfffffff) *)
+ := (Const WO~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
+Notation "'0b00011111111111111111111111111010'" (* 536870906 (0x1ffffffa) *)
+ := (Const WO~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0).
+Notation "'0b00011111111111111111111111111110'" (* 536870910 (0x1ffffffe) *)
+ := (Const WO~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
+Notation "'0b0000000000000111111111111111111111111111111111111111111111111111'" (* 2251799813685247 (0x7ffffffffffffL) *)
+ := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
+Notation "'0b0000000000001111111111111111111111111111111111111111111111011010'" (* 4503599627370458 (0xfffffffffffdaL) *)
+ := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~1~0).
+Notation "'0b0000000000001111111111111111111111111111111111111111111111111110'" (* 4503599627370494 (0xffffffffffffeL) *)
+ := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
diff --git a/src/Compilers/Z/Bounds/Interpretation.v b/src/Compilers/Z/Bounds/Interpretation.v
new file mode 100644
index 000000000..b9880f097
--- /dev/null
+++ b/src/Compilers/Z/Bounds/Interpretation.v
@@ -0,0 +1,177 @@
+Require Import Coq.ZArith.ZArith.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Relations.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Util.ZRange.
+Require Import Crypto.Util.Tactics.DestructHead.
+Export Compilers.Syntax.Notations.
+
+Local Notation eta x := (fst x, snd x).
+Local Notation eta3 x := (eta (fst x), snd x).
+Local Notation eta4 x := (eta3 (fst x), snd x).
+
+Notation bounds := zrange.
+Delimit Scope bounds_scope with bounds.
+Local Open Scope Z_scope.
+
+Module Import Bounds.
+ Definition t := bounds.
+ Bind Scope bounds_scope with t.
+ Local Coercion Z.of_nat : nat >-> Z.
+ Section with_bitwidth.
+ Context (bit_width : option Z).
+ Definition four_corners (f : Z -> Z -> Z) : t -> t -> t
+ := fun x y
+ => let (lx, ux) := x in
+ let (ly, uy) := y in
+ {| lower := Z.min (f lx ly) (Z.min (f lx uy) (Z.min (f ux ly) (f ux uy)));
+ upper := Z.max (f lx ly) (Z.max (f lx uy) (Z.max (f ux ly) (f ux uy))) |}.
+ Definition two_corners (f : Z -> Z) : t -> t
+ := fun x
+ => let (lx, ux) := x in
+ {| lower := Z.min (f lx) (f ux);
+ upper := Z.max (f lx) (f ux) |}.
+ Definition truncation_bounds (b : t)
+ := match bit_width with
+ | Some bit_width => if ((0 <=? lower b) && (upper b <? 2^bit_width))%bool
+ then b
+ else {| lower := 0 ; upper := 2^bit_width - 1 |}
+ | None => b
+ end.
+ Definition BuildTruncated_bounds (l u : Z) : t
+ := truncation_bounds {| lower := l ; upper := u |}.
+ Definition t_map1 (f : Z -> Z) (x : t)
+ := truncation_bounds (two_corners f x).
+ Definition t_map2 (f : Z -> Z -> Z) : t -> t -> t
+ := fun x y => truncation_bounds (four_corners f x y).
+ Definition t_map4 (f : bounds -> bounds -> bounds -> bounds -> bounds) (x y z w : t)
+ := truncation_bounds (f x y z w).
+ Definition add : t -> t -> t := t_map2 Z.add.
+ Definition sub : t -> t -> t := t_map2 Z.sub.
+ Definition mul : t -> t -> t := t_map2 Z.mul.
+ Definition shl : t -> t -> t := t_map2 Z.shiftl.
+ Definition shr : t -> t -> t := t_map2 Z.shiftr.
+ Definition extreme_lor_land_bounds (x y : t) : t
+ := let (lx, ux) := x in
+ let (ly, uy) := y in
+ let lx := Z.abs lx in
+ let ly := Z.abs ly in
+ let ux := Z.abs ux in
+ let uy := Z.abs uy in
+ let max := Z.max (Z.max lx ly) (Z.max ux uy) in
+ {| lower := -2^(1 + Z.log2_up max) ; upper := 2^(1 + Z.log2_up max) |}.
+ Definition extermization_bounds (f : t -> t -> t) (x y : t) : t
+ := truncation_bounds
+ (let (lx, ux) := x in
+ let (ly, uy) := y in
+ if ((lx <? 0) || (ly <? 0))%Z%bool
+ then extreme_lor_land_bounds x y
+ else f x y).
+ Definition land : t -> t -> t
+ := extermization_bounds
+ (fun x y
+ => let (lx, ux) := x in
+ let (ly, uy) := y in
+ {| lower := Z.min 0 (Z.min lx ly) ; upper := Z.max 0 (Z.min ux uy) |}).
+ Definition lor : t -> t -> t
+ := extermization_bounds
+ (fun x y
+ => let (lx, ux) := x in
+ let (ly, uy) := y in
+ {| lower := Z.max lx ly;
+ upper := 2^(Z.max (Z.log2_up (ux+1)) (Z.log2_up (uy+1))) - 1 |}).
+ Definition opp : t -> t := t_map1 Z.opp.
+ Definition neg' (int_width : Z) : t -> t
+ := fun v
+ => let (lb, ub) := v in
+ let might_be_one := ((lb <=? 1) && (1 <=? ub))%Z%bool in
+ let must_be_one := ((lb =? 1) && (ub =? 1))%Z%bool in
+ if must_be_one
+ then {| lower := Z.ones int_width ; upper := Z.ones int_width |}
+ else if might_be_one
+ then {| lower := Z.min 0 (Z.ones int_width) ; upper := Z.max 0 (Z.ones int_width) |}
+ else {| lower := 0 ; upper := 0 |}.
+ Definition neg (int_width : Z) : t -> t
+ := fun v
+ => truncation_bounds (neg' int_width v).
+ Definition cmovne' (r1 r2 : t) : t
+ := let (lr1, ur1) := r1 in
+ let (lr2, ur2) := r2 in
+ {| lower := Z.min lr1 lr2 ; upper := Z.max ur1 ur2 |}.
+ Definition cmovne (x y r1 r2 : t) : t
+ := truncation_bounds (cmovne' r1 r2).
+ Definition cmovle' (r1 r2 : t) : t
+ := let (lr1, ur1) := r1 in
+ let (lr2, ur2) := r2 in
+ {| lower := Z.min lr1 lr2 ; upper := Z.max ur1 ur2 |}.
+ Definition cmovle (x y r1 r2 : t) : t
+ := truncation_bounds (cmovle' r1 r2).
+ End with_bitwidth.
+
+ Module Export Notations.
+ Export Util.ZRange.Notations.
+ Infix "+" := (add _) : bounds_scope.
+ Infix "-" := (sub _) : bounds_scope.
+ Infix "*" := (mul _) : bounds_scope.
+ Infix "<<" := (shl _) : bounds_scope.
+ Infix ">>" := (shr _) : bounds_scope.
+ Infix "&'" := (land _) : bounds_scope.
+ Notation "- x" := (opp _ x) : bounds_scope.
+ End Notations.
+
+ Definition interp_base_type (ty : base_type) : Set := t.
+
+ Definition bit_width_of_base_type ty : option Z
+ := match ty with
+ | TZ => None
+ | TWord logsz => Some (2^Z.of_nat logsz)%Z
+ end.
+
+ Definition interp_op {src dst} (f : op src dst) : interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst
+ := match f in op src dst return interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst with
+ | OpConst T v => fun _ => BuildTruncated_bounds (bit_width_of_base_type T) v v
+ | Add _ _ T => fun xy => add (bit_width_of_base_type T) (fst xy) (snd xy)
+ | Sub _ _ T => fun xy => sub (bit_width_of_base_type T) (fst xy) (snd xy)
+ | Mul _ _ T => fun xy => mul (bit_width_of_base_type T) (fst xy) (snd xy)
+ | Shl _ _ T => fun xy => shl (bit_width_of_base_type T) (fst xy) (snd xy)
+ | Shr _ _ T => fun xy => shr (bit_width_of_base_type T) (fst xy) (snd xy)
+ | Land _ _ T => fun xy => land (bit_width_of_base_type T) (fst xy) (snd xy)
+ | Lor _ _ T => fun xy => lor (bit_width_of_base_type T) (fst xy) (snd xy)
+ | Opp _ T => fun x => opp (bit_width_of_base_type T) x
+ end%bounds.
+
+ Definition of_Z (z : Z) : t := ZToZRange z.
+
+ Definition of_interp t (z : Syntax.interp_base_type t) : interp_base_type t
+ := ZToZRange (interpToZ z).
+
+ Definition bounds_to_base_type (b : t) : base_type
+ := if (0 <=? lower b)%Z
+ then TWord (Z.to_nat (Z.log2_up (Z.log2_up (1 + upper b))))
+ else TZ.
+
+ Definition ComputeBounds {t} (e : Expr base_type op t)
+ (input_bounds : interp_flat_type interp_base_type (domain t))
+ : interp_flat_type interp_base_type (codomain t)
+ := Interp (@interp_op) e input_bounds.
+
+ Definition is_tighter_thanb' {T} : interp_base_type T -> interp_base_type T -> bool
+ := is_tighter_than_bool.
+
+ Definition is_bounded_by' {T} : interp_base_type T -> Syntax.interp_base_type T -> Prop
+ := fun bounds val => is_bounded_by' (bit_width_of_base_type T) bounds (interpToZ val).
+
+ Definition is_tighter_thanb {T} : interp_flat_type interp_base_type T -> interp_flat_type interp_base_type T -> bool
+ := interp_flat_type_relb_pointwise (@is_tighter_thanb').
+
+ Definition is_bounded_by {T} : interp_flat_type interp_base_type T -> interp_flat_type Syntax.interp_base_type T -> Prop
+ := interp_flat_type_rel_pointwise (@is_bounded_by').
+
+ Local Arguments interp_base_type !_ / .
+ Global Instance dec_eq_interp_flat_type {T} : DecidableRel (@eq (interp_flat_type interp_base_type T)) | 10.
+ Proof.
+ induction T; destruct_head base_type; simpl; exact _.
+ Defined.
+End Bounds.
diff --git a/src/Compilers/Z/Bounds/InterpretationLemmas.v b/src/Compilers/Z/Bounds/InterpretationLemmas.v
new file mode 100644
index 000000000..7a1c2bc73
--- /dev/null
+++ b/src/Compilers/Z/Bounds/InterpretationLemmas.v
@@ -0,0 +1,433 @@
+Require Import Coq.ZArith.ZArith.
+Require Import Coq.micromega.Psatz.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.Util.
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Z.Bounds.Interpretation.
+Require Import Crypto.Compilers.SmartMap.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.FixedWordSizesEquality.
+Require Import Crypto.Util.Option.
+Require Import Crypto.Util.Tactics.BreakMatch.
+Require Import Crypto.Util.Tactics.DestructHead.
+Require Import Crypto.Util.Tactics.SpecializeBy.
+Require Import Crypto.Util.Tactics.SplitInContext.
+Require Import Crypto.Util.Tactics.UniquePose.
+
+Local Notation pick_typeb := Bounds.bounds_to_base_type (only parsing).
+Local Notation pick_type v := (SmartFlatTypeMap (fun _ => pick_typeb) v).
+
+Local Open Scope Z_scope.
+
+Local Ltac break_t_step :=
+ first [ progress destruct_head'_and
+ | progress destruct_head'_or
+ | progress destruct_head'_prod
+ | progress split_andb
+ | break_innermost_match_step ].
+
+Local Ltac fin_t :=
+ first [ reflexivity
+ | assumption
+ | match goal with
+ | [ |- and _ _ ]
+ => first [ split; [ | solve [ fin_t ] ]
+ | split; [ solve [ fin_t ] | ] ];
+ try solve [ fin_t ]
+ end
+ | omega ].
+
+Local Ltac specializer_t_step :=
+ first [ progress specialize_by_assumption
+ | progress specialize_by fin_t ].
+
+Local Ltac Zarith_t_step :=
+ first [ match goal with
+ | [ H : (?x <= ?y)%Z, H' : (?y <= ?x)%Z |- _ ]
+ => assert (x = y) by omega; clear H H'
+ end
+ | progress Z.ltb_to_lt_in_context ].
+Local Ltac Zarith_land_lor_t_step :=
+ match goal with
+ | [ |- _ <= Z.lor _ _ <= _ ]
+ => split; etransitivity; [ | apply Z.lor_bounds; omega | apply Z.lor_bounds; omega | ]
+ | [ |- 2^Z.log2_up (?x + 1) - 1 <= 2^Z.log2_up (?y + 1) - 1 ]
+ => let H := fresh in assert (H : x <= y) by omega; rewrite H; reflexivity
+ end.
+Local Ltac word_arith_t :=
+ match goal with
+ | [ |- (0 <= FixedWordSizes.wordToZ ?w <= 2^2^Z.of_nat ?logsz - 1)%Z ]
+ => clear; pose proof (@wordToZ_range logsz w); autorewrite with push_Zof_nat zsimplify_const in *; try omega
+ end.
+
+Local Ltac revert_min_max :=
+ repeat match goal with
+ | [ H : context[Z.min _ _] |- _ ] => revert H
+ | [ H : context[Z.max _ _] |- _ ] => revert H
+ end.
+Local Ltac split_min_max :=
+ repeat match goal with
+ | [ H : (?a <= ?b)%Z |- context[Z.max ?a ?b] ]
+ => rewrite (Z.max_r a b) by omega
+ | [ H : (?b <= ?a)%Z |- context[Z.max ?a ?b] ]
+ => rewrite (Z.max_l a b) by omega
+ | [ H : (?a <= ?b)%Z |- context[Z.min ?a ?b] ]
+ => rewrite (Z.min_l a b) by omega
+ | [ H : (?b <= ?a)%Z |- context[Z.min ?a ?b] ]
+ => rewrite (Z.min_r a b) by omega
+ | [ |- context[Z.max ?a ?b] ]
+ => first [ rewrite (Z.max_l a b) by omega
+ | rewrite (Z.max_r a b) by omega ]
+ | [ |- context[Z.min ?a ?b] ]
+ => first [ rewrite (Z.min_l a b) by omega
+ | rewrite (Z.min_r a b) by omega ]
+ | _ => revert_min_max; progress repeat apply Z.min_case_strong; intros
+ | _ => revert_min_max; progress repeat apply Z.max_case_strong; intros
+ end.
+
+Local Ltac case_Zvar_nonneg_on x :=
+ is_var x;
+ lazymatch type of x with
+ | Z => lazymatch goal with
+ | [ H : (0 <= x)%Z |- _ ] => fail
+ | [ H : (x < 0)%Z |- _ ] => fail
+ | _ => destruct (Z_lt_le_dec x 0); try omega
+ end
+ end.
+Local Ltac case_Zvar_nonneg_step :=
+ match goal with
+ | [ |- context[?x] ]
+ => case_Zvar_nonneg_on x
+ end.
+Local Ltac case_Zvar_nonneg := repeat case_Zvar_nonneg_step.
+
+Local Ltac remove_binary_operation_le_hyps_step :=
+ match goal with
+ | [ H : (?f ?x ?y <= ?f ?x ?y')%Z |- _ ]
+ => assert ((y = y') \/ (y < y' /\ 0 <= x))%Z by (assert (y <= y')%Z by omega; nia);
+ clear H
+ | [ H : (?f ?y ?x <= ?f ?y' ?x)%Z |- _ ]
+ => assert ((y = y') \/ (y < y' /\ 0 <= x))%Z by (assert (y <= y')%Z by omega; nia);
+ clear H
+ | [ H : (?f ?x ?y <= ?f ?x ?y')%Z |- _ ]
+ => assert ((y = y') \/ (y' < y /\ x <= 0))%Z by (assert (y' <= y)%Z by omega; nia);
+ clear H
+ | [ H : (?f ?y ?x <= ?f ?y' ?x)%Z |- _ ]
+ => assert ((y = y') \/ (y' < y /\ x <= 0))%Z by (assert (y' <= y)%Z by omega; nia);
+ clear H
+ | [ H : ?T, H' : ?T |- _ ] => clear H'
+ | [ H : ?A \/ (~?A /\ ?B), H' : ?A \/ (~?A /\ ?C) |- _ ]
+ => assert (A \/ (~A /\ (B /\ C))) by (clear -H H'; tauto); clear H H'
+ | _ => progress destruct_head' or; destruct_head' and; subst; try omega
+ | [ |- (_ <= _ <= _)%Z ] => split
+ | _ => case_Zvar_nonneg_step
+ end.
+
+Local Ltac saturate_with_shift_facts :=
+ repeat match goal with
+ | [ H : ?x <= ?y, H' : ?x' <= ?y' |- context[?x << ?x'] ]
+ => unique assert (x << x' <= y << y') by (apply Z.shiftl_le_mono; omega)
+ | [ H : ?x <= ?y, H' : ?x' <= ?y' |- context[?y << ?y'] ]
+ => unique assert (x << x' <= y << y') by (apply Z.shiftl_le_mono; omega)
+ | [ H : ?x <= ?y, H' : ?x' <= ?y' |- context[?x >> ?x'] ]
+ => unique assert (x >> x' <= y >> y') by (apply Z.shiftr_le_mono; omega)
+ | [ H : ?x <= ?y, H' : ?x' <= ?y' |- context[?y >> ?y'] ]
+ => unique assert (x >> x' <= y >> y') by (apply Z.shiftr_le_mono; omega)
+ end.
+Local Ltac saturate_with_all_shift_facts :=
+ repeat match goal with
+ | _ => progress saturate_with_shift_facts
+ | [ H : ?x <= ?y, H' : ?x' <= ?y' |- context[Z.shiftl _ _] ]
+ => unique assert (x << x' <= y << y') by (apply Z.shiftl_le_mono; omega)
+ | [ H : ?x <= ?y, H' : ?x' <= ?y' |- context[Z.shiftr _ _] ]
+ => unique assert (x >> x' <= y >> y') by (apply Z.shiftr_le_mono; omega)
+ end.
+Local Ltac saturate_land_lor_facts :=
+ repeat match goal with
+ | [ |- context[Z.land ?x ?y] ]
+ => let H := fresh in
+ let H' := fresh in
+ assert (H : 0 <= x) by omega;
+ assert (H' : 0 <= y) by omega;
+ unique pose proof (Z.land_upper_bound_r x y H H');
+ unique pose proof (Z.land_upper_bound_l x y H H')
+ | [ |- context[Z.land ?x ?y] ]
+ => unique assert (0 <= Z.land x y) by (apply Z.land_nonneg; omega)
+ | [ |- context[Z.land ?x ?y] ]
+ => case_Zvar_nonneg_on x; case_Zvar_nonneg_on y
+ | [ |- context[Z.lor ?x ?y] ]
+ => let H := fresh in
+ let H' := fresh in
+ assert (H : 0 <= x) by omega;
+ assert (H' : 0 <= y) by omega;
+ unique pose proof (proj1 (Z.lor_bounds x y H H'));
+ unique pose proof (proj2 (Z.lor_bounds x y H H'))
+ | [ |- context[Z.lor ?x ?y] ]
+ => unique assert (0 <= Z.lor x y) by (apply Z.lor_nonneg; omega)
+ | [ |- context[Z.lor ?x ?y] ]
+ => case_Zvar_nonneg_on x; case_Zvar_nonneg_on y
+ end.
+Local Ltac clean_neg :=
+ repeat match goal with
+ | [ H : (-?x) < 0 |- _ ] => assert (0 <= x) by omega; assert (x <> 0) by omega; clear H
+ | [ H : -?x <= -?y |- _ ] => apply Z.opp_le_mono in H
+ | [ |- -?x <= -?y ] => apply Z.opp_le_mono
+ | _ => progress rewrite <- Z.opp_le_mono in *
+ end.
+Local Ltac replace_with_neg x :=
+ assert (x = -(-x)) by omega; generalize dependent (-x);
+ let x' := fresh in
+ rename x into x'; intro x; intros; subst x';
+ clean_neg.
+Local Ltac replace_all_neg_with_pos :=
+ repeat match goal with
+ | [ H : ?x < 0 |- _ ] => replace_with_neg x
+ end.
+Local Ltac handle_shift_neg :=
+ repeat first [ rewrite !Z.shiftr_opp_r
+ | rewrite !Z.shiftl_opp_r ].
+
+Local Ltac handle_four_corners_step_fast :=
+ first [ progress destruct_head Bounds.t
+ | progress cbv [Bounds.four_corners] in *
+ | progress subst
+ | Zarith_t_step
+ | progress split_min_max
+ | omega
+ | nia ].
+Local Ltac handle_four_corners_step :=
+ first [ handle_four_corners_step_fast
+ | remove_binary_operation_le_hyps_step ].
+Local Ltac handle_four_corners :=
+ lazymatch goal with
+ | [ |- (ZRange.lower (Bounds.four_corners _ _ _) <= _ <= _)%Z ]
+ => idtac
+ end;
+ repeat handle_four_corners_step.
+
+Local Ltac rewriter_t :=
+ first [ rewrite !Bool.andb_true_iff
+ | rewrite !Bool.andb_false_iff
+ | rewrite !Bool.orb_true_iff
+ | rewrite !Bool.orb_false_iff
+ | rewrite !Z.abs_opp
+ | rewrite !Z.max_log2_up
+ | rewrite !Z.add_max_distr_r
+ | rewrite !Z.add_max_distr_l
+ | rewrite wordToZ_ZToWord by (autorewrite with push_Zof_nat zsimplify_const; omega)
+ | match goal with
+ | [ H : _ |- _ ]
+ => first [ rewrite !Bool.andb_true_iff in H
+ | rewrite !Bool.andb_false_iff in H
+ | rewrite !Bool.orb_true_iff in H
+ | rewrite !Bool.orb_false_iff in H ]
+ end ].
+
+Local Arguments Bounds.is_bounded_by' !_ _ _ / .
+
+Lemma is_bounded_by_truncation_bounds Tout bs v
+ (H : Bounds.is_bounded_by (T:=Tbase TZ) bs v)
+ : Bounds.is_bounded_by (T:=Tbase Tout)
+ (Bounds.truncation_bounds (Bounds.bit_width_of_base_type Tout) bs)
+ (ZToInterp v).
+Proof.
+ destruct bs as [l u]; cbv [Bounds.truncation_bounds Bounds.is_bounded_by Bounds.is_bounded_by' Bounds.bit_width_of_base_type ZRange.is_bounded_by'] in *; simpl in *.
+ repeat first [ break_t_step
+ | fin_t
+ | progress simpl in *
+ | Zarith_t_step
+ | rewriter_t
+ | word_arith_t ].
+Qed.
+
+Local Arguments Z.pow : simpl never.
+Local Arguments Z.add !_ !_.
+Local Existing Instances Z.add_le_Proper Z.log2_up_le_Proper Z.pow_Zpos_le_Proper Z.sub_le_eq_Proper.
+Lemma is_bounded_by_interp_op t tR (opc : op t tR)
+ (bs : interp_flat_type Bounds.interp_base_type _)
+ (v : interp_flat_type interp_base_type _)
+ (H : Bounds.is_bounded_by bs v)
+ : Bounds.is_bounded_by (Bounds.interp_op opc bs) (Syntax.interp_op _ _ opc v).
+Proof.
+ destruct opc; apply is_bounded_by_truncation_bounds;
+ repeat first [ progress simpl in *
+ | progress cbv [interp_op lift_op cast_const Bounds.interp_base_type Bounds.is_bounded_by' ZRange.is_bounded_by'] in *
+ | progress destruct_head'_prod
+ | progress destruct_head'_and
+ | omega
+ | match goal with
+ | [ |- context[interpToZ ?x] ]
+ => generalize dependent (interpToZ x); clear x; intros
+ | [ |- _ /\ True ] => split; [ | tauto ]
+ end ].
+ { handle_four_corners. }
+ { handle_four_corners. }
+ { handle_four_corners. }
+ { destruct_head Bounds.t.
+ case_Zvar_nonneg; replace_all_neg_with_pos; handle_shift_neg;
+ autorewrite with Zshift_to_pow;
+ rewrite ?Z.div_opp_l_complete by auto with zarith;
+ autorewrite with Zpow_to_shift.
+ 16:split_min_max; saturate_with_shift_facts; omega.
+ all:admit. }
+ { destruct_head Bounds.t.
+ case_Zvar_nonneg; replace_all_neg_with_pos; handle_shift_neg; admit. }
+ { repeat first [ progress destruct_head Bounds.t
+ | progress simpl in *
+ | break_t_step
+ | Zarith_t_step
+ | rewriter_t
+ | progress replace_all_neg_with_pos
+ | progress saturate_land_lor_facts
+ | split_min_max; omega ];
+ admit. }
+ { repeat first [ progress destruct_head Bounds.t
+ | progress simpl in *
+ | break_t_step
+ | Zarith_t_step
+ | rewriter_t
+ | progress replace_all_neg_with_pos
+ | progress saturate_land_lor_facts
+ | progress Zarith_land_lor_t_step
+ | solve [ split_min_max; try omega; try Zarith_land_lor_t_step ] ];
+ admit. }
+ { repeat first [ progress destruct_head Bounds.t
+ | progress simpl in *
+ | progress split_min_max
+ | omega ]. }
+Admitted.
+
+Local Arguments lift_op : simpl never.
+Local Arguments cast_back_flat_const : simpl never.
+Local Arguments unzify_op : simpl never.
+Local Arguments Z.pow : simpl never.
+Local Arguments Z.add !_ !_.
+Local Existing Instance Z.pow_Zpos_le_Proper.
+Lemma pull_cast_genericize_op t tR (opc : op t tR)
+ (bs : interp_flat_type Bounds.interp_base_type t)
+ (v : interp_flat_type interp_base_type (pick_type bs))
+ (H : Bounds.is_bounded_by bs
+ (SmartFlatTypeMapUnInterp
+ (fun (t1 : base_type) (b0 : Bounds.interp_base_type t1) => cast_const) v))
+ : interp_op t tR opc (cast_back_flat_const v)
+ = cast_back_flat_const (interp_op (pick_type bs) (pick_type (Bounds.interp_op opc bs)) (genericize_op opc) v).
+Proof.
+ pose proof (is_bounded_by_interp_op t tR opc bs).
+ unfold interp_op in *.
+ rewrite Zinterp_op_genericize_op.
+ generalize dependent (Zinterp_op t tR opc).
+ generalize dependent (Bounds.interp_op opc bs); clear opc; simpl; intros.
+ revert dependent t; induction tR as [tR| |]; intros;
+ [
+ | repeat first [ match goal with
+ | [ |- ?x = ?y ]
+ => transitivity tt; destruct x, y; reflexivity
+ end
+ | reflexivity
+ | progress simpl @Bounds.is_bounded_by in *
+ | rewrite !lift_op_prod_dst
+ | rewrite !cast_back_flat_const_prod
+ | progress split_and
+ | match goal with
+ | [ H : _ |- _ ] => first [ setoid_rewrite lift_op_prod_dst in H
+ | setoid_rewrite cast_back_flat_const_prod in H ]
+ end
+ | setoid_rewrite lift_op_prod_dst
+ | match goal with
+ | [ H : _ |- _ ] => erewrite H by eassumption
+ end ].. ].
+ revert dependent tR; induction t as [t| |]; intros;
+ [
+ | repeat first [ match goal with
+ | [ |- ?x = ?y ]
+ => transitivity tt; destruct x, y; reflexivity
+ end
+ | reflexivity
+ | progress simpl @Bounds.is_bounded_by in *
+ | rewrite !lift_op_prod_dst
+ | rewrite !cast_back_flat_const_prod
+ | progress split_and
+ | match goal with
+ | [ H : _ |- _ ] => first [ setoid_rewrite lift_op_prod_dst in H
+ | setoid_rewrite cast_back_flat_const_prod in H ]
+ end
+ | setoid_rewrite lift_op_prod_dst
+ | match goal with
+ | [ H : _ |- _ ] => erewrite H by eassumption
+ end ].. ].
+ { simpl in *; unfold unzify_op, cast_back_flat_const, SmartFlatTypeMap, Bounds.interp_base_type, cast_const, Bounds.is_bounded_by', lift_op, SmartFlatTypeMapUnInterp, SmartFlatTypeMapInterp2, cast_const in *; simpl in *.
+ unfold Bounds.is_bounded_by', cast_const, ZToInterp, interpToZ, Bounds.bounds_to_base_type, ZRange.is_bounded_by' in *; simpl in *.
+ destruct_head base_type; break_innermost_match; Z.ltb_to_lt; destruct_head Bounds.t;
+ repeat match goal with
+ | _ => progress destruct_head'_and
+ | _ => reflexivity
+ | [ H : forall v, _ /\ True -> _ |- _ ] => specialize (fun v pf => H v (conj pf I))
+ | [ H : forall v, _ -> _ /\ True |- _ ] => pose proof (fun v pf => proj1 (H v pf)); clear H
+ | [ H : True |- _ ] => clear H
+ | [ H : ?T, H' : ?T |- _ ] => clear H
+ | [ H : forall v, _ -> _ <= ?f v <= _ |- ?f ?v' = _ ]
+ => specialize (H v')
+ | [ H : forall v, _ -> _ <= ?f (?g v) <= _ |- ?f (?g ?v') = _ ]
+ => specialize (H v')
+ | [ H : forall v, _ -> _ <= ?f (?g (?h v)) <= _ /\ _ /\ _ |- context[?h ?v'] ]
+ => specialize (H v')
+ | [ H : forall v, _ -> _ <= ?f (?g (?h (?i v))) <= _ /\ _ /\ _ |- context[?h (?i ?v')] ]
+ => specialize (H v')
+ | _ => progress specialize_by omega
+ | _ => rewrite wordToZ_ZToWord
+ by repeat match goal with
+ | [ |- and _ _ ] => split
+ | [ |- ?x < ?y ] => cut (1 + x <= y); [ omega | ]
+ | _ => omega
+ | _ => progress autorewrite with push_Zof_nat zsimplify_const
+ | _ => rewrite Z2Nat.id by auto with zarith
+ | _ => rewrite <- !Z.log2_up_le_full
+ end
+ | _ => rewrite wordToZ_ZToWord in *
+ by repeat match goal with
+ | [ |- and _ _ ] => split
+ | [ |- ?x < ?y ] => cut (1 + x <= y); [ omega | ]
+ | _ => omega
+ | _ => progress autorewrite with push_Zof_nat zsimplify_const
+ | _ => rewrite Z2Nat.id by auto with zarith
+ | _ => rewrite <- !Z.log2_up_le_full
+ end
+ | _ => rewrite wordToZ_ZToWord_wordToZ
+ by (rewrite Nat2Z.inj_le, Z2Nat.id, <- !Z.log2_up_le_pow2_full by auto with zarith; omega)
+ | _ => rewrite wordToZ_ZToWord_wordToZ in *
+ by (rewrite Nat2Z.inj_le, Z2Nat.id, <- !Z.log2_up_le_pow2_full by auto with zarith; omega)
+ end.
+ all:admit. }
+ { simpl in *.
+ specialize (H0 tt I).
+ simpl in *.
+ hnf in H0.
+ unfold cast_back_flat_const, lift_op, unzify_op in *; simpl in *.
+ unfold interpToZ in *.
+ unfold Bounds.bounds_to_base_type in *.
+ destruct_head base_type; simpl in *.
+ split_andb.
+ Z.ltb_to_lt.
+ all:destruct_head' and.
+ all:simpl in *.
+ all:break_innermost_match; break_match_hyps; split_andb; Z.ltb_to_lt; try reflexivity.
+ all:try (simpl in *;
+ rewrite wordToZ_ZToWord
+ by (autorewrite with push_Zof_nat zsimplify_const;
+ rewrite Z2Nat.id by auto with zarith;
+ split; try omega;
+ match goal with
+ | [ |- (?x < ?y)%Z ]
+ => apply (Z.lt_le_trans x (x + 1) y); [ omega | ]
+ end;
+ rewrite <- !Z.log2_up_le_full;
+ omega)).
+ all:try reflexivity.
+ unfold interpToZ, cast_const.
+ simpl.
+ rewrite ZToWord_wordToZ_ZToWord; [ reflexivity | ].
+ apply Nat2Z.inj_le.
+ rewrite Z2Nat.id by auto with zarith.
+
+Admitted.
diff --git a/src/Compilers/Z/Bounds/MapCastByDeBruijn.v b/src/Compilers/Z/Bounds/MapCastByDeBruijn.v
new file mode 100644
index 000000000..45b084566
--- /dev/null
+++ b/src/Compilers/Z/Bounds/MapCastByDeBruijn.v
@@ -0,0 +1,23 @@
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Z.MapCastByDeBruijn.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.Util.
+Require Import Crypto.Compilers.Z.Bounds.Interpretation.
+
+Section language.
+ Context {t : type base_type}.
+
+ Definition MapCastCompile := @MapCastCompile t.
+ Definition MapCastDoCast
+ := @MapCastDoCast
+ (@Bounds.interp_base_type) (@Bounds.interp_op)
+ (fun _ => @Bounds.bounds_to_base_type)
+ (fun _ _ opc _ => @genericize_op _ _ _ opc _ _ _)
+ t.
+ Definition MapCastDoInterp
+ := @MapCastDoInterp
+ (@Bounds.interp_base_type) (fun _ => @Bounds.bounds_to_base_type)
+ t.
+ Definition MapCast e input_bounds
+ := MapCastDoInterp input_bounds (MapCastDoCast input_bounds (MapCastCompile e)).
+End language.
diff --git a/src/Compilers/Z/Bounds/MapCastByDeBruijnInterp.v b/src/Compilers/Z/Bounds/MapCastByDeBruijnInterp.v
new file mode 100644
index 000000000..46f472311
--- /dev/null
+++ b/src/Compilers/Z/Bounds/MapCastByDeBruijnInterp.v
@@ -0,0 +1,25 @@
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Wf.
+Require Import Crypto.Compilers.SmartMap.
+Require Import Crypto.Compilers.Relations.
+Require Import Crypto.Compilers.Z.MapCastByDeBruijnInterp.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.Util.
+Require Import Crypto.Compilers.Z.Bounds.Interpretation.
+Require Import Crypto.Compilers.Z.Bounds.InterpretationLemmas.
+Require Import Crypto.Compilers.Z.Bounds.MapCastByDeBruijn.
+
+Lemma MapCastCorrect
+ {t} (e : Expr base_type op t)
+ (Hwf : Wf e)
+ (input_bounds : interp_flat_type Bounds.interp_base_type (domain t))
+ : forall {b} e' (He':MapCast e input_bounds = Some (existT _ b e'))
+ v v' (Hv : Bounds.is_bounded_by input_bounds v /\ cast_back_flat_const v' = v),
+ Interp (@Bounds.interp_op) e input_bounds = b
+ /\ Bounds.is_bounded_by b (Interp interp_op e v)
+ /\ cast_back_flat_const (Interp interp_op e' v') = (Interp interp_op e v).
+Proof.
+ apply MapCastCorrect; auto.
+ { apply is_bounded_by_interp_op. }
+ { apply pull_cast_genericize_op. }
+Qed.
diff --git a/src/Compilers/Z/Bounds/MapCastByDeBruijnWf.v b/src/Compilers/Z/Bounds/MapCastByDeBruijnWf.v
new file mode 100644
index 000000000..04d1f964e
--- /dev/null
+++ b/src/Compilers/Z/Bounds/MapCastByDeBruijnWf.v
@@ -0,0 +1,41 @@
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Wf.
+Require Import Crypto.Compilers.SmartMap.
+Require Import Crypto.Compilers.Relations.
+Require Import Crypto.Compilers.Z.MapCastByDeBruijnWf.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.Util.
+Require Import Crypto.Compilers.Z.Bounds.Interpretation.
+Require Import Crypto.Compilers.Z.Bounds.InterpretationLemmas.
+Require Import Crypto.Compilers.Z.Bounds.MapCastByDeBruijn.
+
+Definition Wf_MapCast
+ {t} (e : Expr base_type op t)
+ (input_bounds : interp_flat_type Bounds.interp_base_type (domain t))
+ {b} e' (He' : MapCast e input_bounds = Some (existT _ b e'))
+ (Hwf : Wf e)
+ : Wf e'
+ := @Wf_MapCast
+ (@Bounds.interp_base_type) (@Bounds.interp_op)
+ (fun _ => @Bounds.bounds_to_base_type)
+ (fun _ _ opc _ => @genericize_op _ _ _ opc _ _ _)
+ t e input_bounds b e' He' Hwf.
+
+Definition Wf_MapCast_arrow
+ {s d} (e : Expr base_type op (Arrow s d))
+ (input_bounds : interp_flat_type Bounds.interp_base_type s)
+ {b} e' (He' : MapCast e input_bounds = Some (existT _ b e'))
+ (Hwf : Wf e)
+ : Wf e'
+ := @Wf_MapCast_arrow
+ (@Bounds.interp_base_type) (@Bounds.interp_op)
+ (fun _ => @Bounds.bounds_to_base_type)
+ (fun _ _ opc _ => @genericize_op _ _ _ opc _ _ _)
+ s d e input_bounds b e' He' Hwf.
+
+Hint Extern 1 (Wf ?e')
+=> match goal with
+ | [ He : MapCast _ _ = Some (existT _ _ e') |- _ ]
+ => first [ refine (@Wf_MapCast _ _ _ _ _ He _)
+ | refine (@Wf_MapCast_arrow _ _ _ _ _ _ He _) ]
+ end : wf.
diff --git a/src/Compilers/Z/Bounds/Pipeline.v b/src/Compilers/Z/Bounds/Pipeline.v
new file mode 100644
index 000000000..6c9cda840
--- /dev/null
+++ b/src/Compilers/Z/Bounds/Pipeline.v
@@ -0,0 +1,20 @@
+(** * Reflective Pipeline *)
+Require Import Crypto.Compilers.Z.Bounds.Pipeline.Glue.
+Require Import Crypto.Compilers.Z.Bounds.Pipeline.ReflectiveTactics.
+(** This file combines the various PHOAS modules in tactics,
+ culminating in a tactic [refine_reflectively], which solves a goal of the form
+<<
+cast_back_flat_const (?x args) = f (cast_back_flat_const args)
+ /\ Bounds.is_bounded_by ?bounds (?x args)
+>>
+while instantiating [?x] and [?bounds] with nicely-reduced constants.
+ *)
+
+Module Export Exports.
+ Export Glue.Exports.
+ Export ReflectiveTactics.Exports.
+End Exports.
+
+Ltac refine_reflectively :=
+ refine_to_reflective_glue;
+ do_reflective_pipeline.
diff --git a/src/Compilers/Z/Bounds/Pipeline/Definition.v b/src/Compilers/Z/Bounds/Pipeline/Definition.v
new file mode 100644
index 000000000..ae3401b78
--- /dev/null
+++ b/src/Compilers/Z/Bounds/Pipeline/Definition.v
@@ -0,0 +1,177 @@
+(** * Reflective Pipeline: Main Pipeline Definition *)
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Bounds.Pipeline.OutputType.
+(** This file contains the definitions of the assembling of the
+ various transformations that are used in the pipeline. There are
+ two stages to the reflective pipeline, with different
+ requirements.
+
+ The pre-Wf stage is intended to consist of transformations that
+ make the term smaller, and, importantly, should only consist of
+ transformations whose interpretation-correctness proofs do not
+ require well-founded hypotheses. Generally this is the case for
+ transformations whose input and output [var] types match. The
+ correctness condition for this stage is that the interpretation of
+ the transformed term must equal the interpretation of the original
+ term, with no side-conditions.
+
+ The post-Wf stage is the rest of the pipeline; its correctness
+ condition must have the shape of the correctness condition for
+ word-size selection. We define a record to hold the transformed
+ term, so that we can get bounds and similar out of it, without
+ running into issues with slowness of conversion. *)
+
+(** ** Pre-Wf Stage *)
+(** *** Pre-Wf Pipeline Imports *)
+Require Import Crypto.Compilers.Eta.
+Require Import Crypto.Compilers.EtaInterp.
+Require Import Crypto.Compilers.Z.ArithmeticSimplifier.
+Require Import Crypto.Compilers.Z.ArithmeticSimplifierInterp.
+
+(** *** Definition of the Pre-Wf Pipeline *)
+(** Do not change the name or the type of this definition *)
+Definition PreWfPipeline {t} (e : Expr base_type op t) : Expr base_type op _
+ := ExprEta (SimplifyArith e).
+
+(** *** Correctness proof of the Pre-Wf Pipeline *)
+(** Do not change the statement of this lemma. You shouldn't need to
+ change it's proof, either; all of the relevant lemmas should be in
+ the [reflective_interp] rewrite database. If they're not, you
+ should find the file where they are defined and add them. *)
+Lemma InterpPreWfPipeline {t} (e : Expr base_type op t)
+ : forall x, Interp interp_op (PreWfPipeline e) x = Interp interp_op e x.
+Proof.
+ unfold PreWfPipeline; intro.
+ repeat autorewrite with reflective_interp.
+ reflexivity.
+Qed.
+
+
+
+(** ** Post-Wf Stage *)
+(** *** Post-Wf Pipeline Imports *)
+Require Import Crypto.Compilers.Z.Bounds.Interpretation.
+Require Import Crypto.Compilers.EtaWf.
+Require Import Crypto.Compilers.Z.Inline.
+Require Import Crypto.Compilers.Z.InlineInterp.
+Require Import Crypto.Compilers.Z.InlineWf.
+Require Import Crypto.Compilers.Linearize.
+Require Import Crypto.Compilers.LinearizeInterp.
+Require Import Crypto.Compilers.LinearizeWf.
+Require Import Crypto.Compilers.Z.Bounds.MapCastByDeBruijn.
+Require Import Crypto.Compilers.Z.Bounds.MapCastByDeBruijnInterp.
+Require Import Crypto.Compilers.Z.Bounds.MapCastByDeBruijnWf.
+Require Import Crypto.Util.Sigma.MapProjections.
+
+(** *** Definition of the Post-Wf Pipeline *)
+(** Do not change the name or the type of this definition *)
+Definition PostWfPipeline
+ {t} (e : Expr base_type op t)
+ (input_bounds : interp_flat_type Bounds.interp_base_type (domain t))
+ : option ProcessedReflectivePackage
+ := Build_ProcessedReflectivePackage_from_option_sigma
+ e input_bounds
+ (let e := Linearize e in
+ let e := InlineConst e in
+ let e := MapCast e input_bounds in
+ option_map
+ (projT2_map
+ (fun b e'
+ => let e' := InlineConst e' in
+ let e' := ExprEta e' in
+ e'))
+ e).
+
+(** *** Correctness proof of the Pre-Wf Pipeline *)
+(** Do not change the statement of this lemma. *)
+Require Import Crypto.Compilers.Wf.
+Require Import Crypto.Compilers.Equality.
+Require Import Crypto.Compilers.SmartMap.
+Require Import Crypto.Compilers.Z.Syntax.Util.
+Require Import Crypto.Util.Tactics.BreakMatch.
+Require Import Crypto.Util.Option.
+Require Import Crypto.Util.Sigma.
+Require Import Crypto.Util.Prod.
+Require Import Crypto.Util.HProp.
+Require Import Crypto.Util.Decidable.
+
+Local Notation pick_typeb := Bounds.bounds_to_base_type (only parsing).
+Local Notation pick_type v := (SmartFlatTypeMap (fun _ => pick_typeb) v).
+Definition PostWfPipelineCorrect
+ {t}
+ (e : Expr base_type op t)
+ (input_bounds : interp_flat_type Bounds.interp_base_type (domain t))
+ (Hwf : Wf e)
+ {b e'} (He : PostWfPipeline e input_bounds
+ = Some {| input_expr := e ; input_bounds := input_bounds ; output_bounds := b ; output_expr := e' |})
+ (v : interp_flat_type Syntax.interp_base_type (domain t))
+ (v' : interp_flat_type Syntax.interp_base_type (pick_type input_bounds))
+ (Hv : Bounds.is_bounded_by input_bounds v /\ cast_back_flat_const v' = v)
+ : Interp (@Bounds.interp_op) e input_bounds = b
+ /\ Bounds.is_bounded_by b (Interp interp_op e v)
+ /\ cast_back_flat_const (Interp interp_op e' v') = Interp interp_op e v.
+Proof.
+ (** These first two lines probably shouldn't change much *)
+ unfold PostWfPipeline, Build_ProcessedReflectivePackage_from_option_sigma, option_map, projT2_map in *.
+ repeat (break_match_hyps || inversion_option || inversion_ProcessedReflectivePackage
+ || inversion_sigma || eliminate_hprop_eq || inversion_prod
+ || simpl in * || subst).
+ (** Now handle all the transformations that come after the word-size selection *)
+ rewrite InterpExprEta_arrow, InterpInlineConst
+ by eauto with wf.
+ (** Now handle all the transformations that come before the word-size selection *)
+ rewrite <- !InterpLinearize with (e:=e), <- !(@InterpInlineConst _ _ _ (Linearize e))
+ by eauto with wf.
+ (** Now handle word-size selection itself *)
+ eapply MapCastCorrect; eauto with wf.
+Qed.
+
+
+(** ** Constant Simplification and Unfolding *)
+(** The reflective pipeline may introduce constants that you want to
+ unfold before instantiating the refined term; you can control that
+ here. A number of reflection-specific constants are always
+ unfolded (in ReflectiveTactics.v). Currently, we also reduce
+ expressions of the form [wordToZ (ZToWord Z_literal)], as
+ specified here. *)
+Require Import Coq.ZArith.ZArith.
+Require Import Crypto.Util.FixedWordSizes.
+Require Import Bedrock.Word.
+
+Module Export Exports. (* export unfolding strategy *)
+ (* iota is probably (hopefully?) the cheapest reduction.
+ Unfortunately, we can't say no-op here. This is meant to be
+ extended. *)
+ Declare Reduction extra_interp_red := cbv iota.
+
+ (** Overload this to change reduction behavior of constants of the
+ form [wordToZ (ZToWord Z_literal)]. You might want to set this
+ to false if your term is very large, to speed things up. *)
+ Ltac do_constant_simplification := constr:(true).
+
+ Global Arguments ZToWord !_ !_ / .
+ Global Arguments wordToZ !_ !_ / .
+ Global Arguments word_case_dep _ !_ _ _ _ _ / .
+ Global Arguments ZToWord32 !_ / .
+ Global Arguments ZToWord64 !_ / .
+ Global Arguments ZToWord128 !_ / .
+ Global Arguments ZToWord_gen !_ !_ / .
+ Global Arguments word32ToZ !_ / .
+ Global Arguments word64ToZ !_ / .
+ Global Arguments word128ToZ !_ / .
+ Global Arguments wordToZ_gen !_ !_ / .
+ Global Arguments Z.to_N !_ / .
+ Global Arguments Z.of_N !_ / .
+ Global Arguments Word.NToWord !_ !_ / .
+ Global Arguments Word.wordToN !_ !_ / .
+ Global Arguments Word.posToWord !_ !_ / .
+ Global Arguments N.succ_double !_ / .
+ Global Arguments Word.wzero' !_ / .
+ Global Arguments N.double !_ .
+ Global Arguments Nat.pow !_ !_ / .
+ Global Arguments Nat.mul !_ !_ / .
+ Global Arguments Nat.add !_ !_ / .
+
+ Declare Reduction constant_simplification := cbn [FixedWordSizes.wordToZ FixedWordSizes.ZToWord word_case_dep ZToWord32 ZToWord64 ZToWord128 ZToWord_gen word32ToZ word64ToZ word128ToZ wordToZ_gen Word.NToWord Word.wordToN Word.posToWord Word.wzero' Z.to_N Z.of_N N.succ_double N.double Nat.pow Nat.mul Nat.add].
+End Exports.
diff --git a/src/Compilers/Z/Bounds/Pipeline/Glue.v b/src/Compilers/Z/Bounds/Pipeline/Glue.v
new file mode 100644
index 000000000..5aa22e106
--- /dev/null
+++ b/src/Compilers/Z/Bounds/Pipeline/Glue.v
@@ -0,0 +1,456 @@
+(** * Reflective Pipeline: Glue Code *)
+(** This file defines the tactics that transform a non-reflective goal
+ into a goal the that the reflective machinery can handle. *)
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Reify.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.Util.
+Require Import Crypto.Compilers.Z.Reify.
+Require Import Crypto.Compilers.Z.Bounds.Interpretation.
+Require Import Crypto.Util.Tactics.Head.
+Require Import Crypto.Util.Curry.
+Require Import Crypto.Util.FixedWordSizes.
+Require Import Crypto.Util.BoundedWord.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Sigma.Associativity.
+Require Import Crypto.Util.Sigma.MapProjections.
+Require Import Crypto.Util.Tactics.EvarExists.
+Require Import Crypto.Util.Tactics.GetGoal.
+Require Import Crypto.Util.Tactics.PrintContext.
+Require Import Crypto.Util.Tactics.MoveLetIn.
+
+Module Export Exports.
+ Export Crypto.Compilers.Z.Reify. (* export for the tactic redefinitions *)
+End Exports.
+
+(** ** [reassoc_sig_and_eexists] *)
+(** The [reassoc_sig_and_eexists] tactic operates on a goal convertible with
+<<
+{ f : { a | is_bounded_by bounds a }
+| BoundedWordToZ f = rexprZ (BoundedWordToZ a) ... (BoundedWordToZ z) }
+>>
+ and leaves a goal of the form
+<<
+is_bounded_by bounds (map wordToZ ?f)
+ /\ map wordToZ ?f = rexprZ (map wordToZ (proj1_sig a)) ... (map wordToZ (proj1_sig z))
+>>
+ where [?f] is a context variable set to a new evar. This tactic
+ relies on the exact definition of [BoundedWordToZ]. *)
+
+
+(** The tactic [unfold_paired_tuple_map] unfolds any [Tuple.map]s
+ applied to [pair]s. *)
+Ltac unfold_paired_tuple_map :=
+ repeat match goal with
+ | [ |- context[Tuple.map (n:=S ?N) _ (pair _ _)] ]
+ => progress change (@Tuple.map (S N)) with (fun A B f => @Tuple.map' A B f N); cbv beta iota delta [Tuple.map']
+ end.
+(** The tactic [change_to_reified_type f] reifies the type of a
+ context variable [f] and changes [f] to the interpretation of that
+ type. *)
+Ltac change_to_reified_type f :=
+ let T := type of f in
+ let cT := (eval compute in T) in
+ let rT := reify_type cT in
+ change (interp_type Syntax.interp_base_type rT) in (type of f).
+
+(** The tactic [sig_dlet_in_rhs_to_context_curried] moves to the
+ context any [dlet x := y in ...] on the rhs of a goal of the form
+ [{ a | lhs = rhs }], curries each such moved definition, and then
+ reifies the type of each such context variable. *)
+Ltac sig_dlet_in_rhs_to_context_curried :=
+ lazymatch goal with
+ | [ |- { a | _ = @Let_In ?A ?B ?x _ } ]
+ => let f := fresh in
+ sig_dlet_in_rhs_to_context_step f;
+ change_with_curried f;
+ change_to_reified_type f;
+ sig_dlet_in_rhs_to_context_curried
+ | _ => idtac
+ end.
+(** The tactic [preunfold_and_dlet_to_context] will unfold
+ [BoundedWordToZ] and [Tuple.map]s applied to [pair]s, and then
+ look for a [dlet x := y in ...] in the RHS of a goal of shape [{a
+ | LHS = RHS }] and replace it with a context variable. *)
+Ltac preunfold_and_dlet_to_context :=
+ unfold_paired_tuple_map;
+ cbv [BoundedWordToZ]; cbn [fst snd proj1_sig];
+ sig_dlet_in_rhs_to_context_curried.
+(** The tactic [pattern_proj1_sig_in_lhs_of_sig] takes a goal of the form
+<<
+{ a : A | P }
+>>
+ where [A] is a sigma type, and leaves a goal of the form
+<<
+{ a : A | dlet p := P' in p (proj1_sig a)
+>>
+ where all occurrences of [proj1_sig a] have been abstracted out of
+ [P] to make [P']. *)
+Ltac pattern_proj1_sig_in_sig :=
+ eapply proj2_sig_map;
+ [ let a := fresh in
+ let H := fresh in
+ intros a H; pattern (proj1_sig a);
+ lazymatch goal with
+ | [ |- ?P ?p1a ]
+ => cut (dlet p := P in p p1a);
+ [ clear; abstract (cbv [Let_In]; exact (fun x => x)) | ]
+ end;
+ exact H
+ | cbv beta ].
+(** The tactic [pattern_sig_sig_assoc] takes a goal of the form
+<<
+{ a : { a' : A | P } | Q }
+>>
+ where [Q] mentions [proj1_sig a] but not [proj2_sig a], and leaves
+ a goal of the form
+<<
+{ a : A | P /\ Q }
+>>
+ *)
+Ltac pattern_sig_sig_assoc :=
+ pattern_proj1_sig_in_sig;
+ let f := fresh in
+ goal_dlet_to_context_step f;
+ apply sig_sig_assoc;
+ subst f; cbv beta.
+(** The tactic [reassoc_sig_and_eexists] will unfold [BoundedWordToZ]
+ and move any [dlet x := ... in ...] to context variables, and then
+ take a goal of the form
+<<
+{ a : { a' : A | P a' } | Q (proj1_sig a) }
+>>
+ where [Q] mentions [proj1_sig a] but not [proj2_sig a], and leave
+ a goal of the form
+<<
+P ?a /\ Q ?a
+>>
+ *)
+Ltac reassoc_sig_and_eexists :=
+ preunfold_and_dlet_to_context;
+ pattern_sig_sig_assoc;
+ evar_exists.
+
+
+(** ** [do_curry_rhs] *)
+(** The [do_curry_rhs] tactic takes a goal of the form
+<<
+_ /\ _ = F A B ... Z
+>>
+ and turns it into a goal of the form
+<<
+_ /\ _ = F' (A, B, ..., Z)
+>>
+ *)
+Ltac do_curry_rhs :=
+ lazymatch goal with
+ | [ |- _ /\ _ = ?f_Z ]
+ => let f_Z := head f_Z in
+ change_with_curried f_Z
+ end.
+
+(** ** [split_BoundedWordToZ] *)
+(** The [split_BoundedWordToZ] tactic takes a goal of the form
+<<
+_ /\ (map wordToZ (proj1_sig f1), ... map wordToZ (proj1_sig fn)) = F ARGS
+>>
+ and splits [f1] ... [fn] and any arguments in ARGS into two
+ parts, one part about the computational behavior, and another part
+ about the boundedness.
+
+ This pipeline relies on the specific definition of
+ [BoundedWordToZ], and requires [f] to be a context variable which
+ is set to a single evar. *)
+(** First we ensure the goal has the right shape, and give helpful
+ error messages if it does not *)
+Ltac check_fW_type descr top_fW fW :=
+ lazymatch fW with
+ | fst ?fW => check_fW_type descr top_fW fW
+ | snd ?fW => check_fW_type descr top_fW fW
+ | _ => let G := get_goal in
+ let shape := uconstr:(map wordToZ ?fW) in
+ let efW := uconstr:(?fW) in
+ first [ is_var fW
+ | fail 1 "In the goal" G
+ descr shape
+ "where" efW "must be a repeated application of fst and snd"
+ "to a single context variable which is defined to be an evar."
+ "The term" top_fW "is based on" fW "which is not a variable" ];
+ match goal with
+ | [ fW' := ?e |- _ ]
+ => constr_eq fW' fW;
+ first [ is_evar e
+ | fail 2 "In the goal" G
+ descr shape
+ "where" efW "must be a repeated application of fst and snd"
+ "to a single context variable which is defined to be an evar."
+ "The term" top_fW "is based on" fW' "which is a context variable"
+ "with body" e "which is not a bare evar" ]
+ | [ fW' : _ |- _ ]
+ => constr_eq fW fW';
+ fail 1 "In the goal" G
+ descr shape
+ "where" efW "must be a repeated application of fst and snd"
+ "to a single context variable which is defined to be an evar."
+ "The term" top_fW "is based on" fW' "which is a context variable without a body"
+ | _ => fail 1 "In the goal" G
+ descr shape
+ "where" efW "must be a repeated application of fst and snd"
+ "to a single context variable which is defined to be an evar."
+ "The term" top_fW "is based on" fW "which is not a context variable"
+ end
+ end.
+Tactic Notation "check_fW_type" string(descr) constr(fW)
+ := check_fW_type descr fW fW.
+Ltac check_is_bounded_by_shape subterm_type :=
+ lazymatch subterm_type with
+ | ZRange.is_bounded_by None ?bounds (map wordToZ ?fW)
+ => check_fW_type "The ℤ argument to is_bounded_by must have the shape" fW
+ | ?A /\ ?B
+ => check_is_bounded_by_shape A;
+ check_is_bounded_by_shape B
+ | _ => let G := get_goal in
+ let shape := uconstr:(ZRange.is_bounded_by None ?bounds (map wordToZ ?fW)) in
+ fail "In the goal" G
+ "The first conjunct of the goal is expected to be a conjunction of things of the shape" shape
+ "but a subterm not matching this shape was found:" subterm_type
+ end.
+Ltac check_LHS_Z_shape subterm :=
+ lazymatch subterm with
+ | map wordToZ ?fW
+ => check_fW_type "The left-hand side of the second conjunct of the goal must be a tuple of terms with shape" fW
+ | (?A, ?B)
+ => check_LHS_Z_shape A;
+ check_LHS_Z_shape B
+ | _ => let G := get_goal in
+ let shape := uconstr:(map wordToZ ?fW) in
+ fail "In the goal" G
+ "The second conjunct of the goal is expected to be a equality whose"
+ "left-hand side is a tuple of terms of the shape" shape
+ "but a subterm not matching this shape was found:" subterm
+ end.
+Ltac check_RHS_Z_shape_rec subterm :=
+ lazymatch subterm with
+ | map wordToZ ?fW
+ => idtac
+ | (?A, ?B)
+ => check_RHS_Z_shape_rec A;
+ check_RHS_Z_shape_rec B
+ | _ => let G := get_goal in
+ let shape := uconstr:(map wordToZ ?fW) in
+ fail "In the goal" G
+ "The second conjunct of the goal is expected to be a equality whose"
+ "right-hand side is the application of a function to a tuple of terms of the shape" shape
+ "but a subterm not matching this shape was found:" subterm
+ end.
+Ltac check_RHS_Z_shape RHS :=
+ lazymatch RHS with
+ | ?f ?args
+ => let G := get_goal in
+ first [ is_var f
+ | fail 1 "In the goal" G
+ "The second conjunct of the goal is expected to be a equality whose"
+ "right-hand side is the application of a single context-variable to a tuple"
+ "but the right-hand side is" RHS
+ "which is an application of something which is not a context variable:" f ];
+ check_RHS_Z_shape_rec args
+ | _ => let G := get_goal in
+ let shape := uconstr:(map wordToZ ?fW) in
+ fail "In the goal" G
+ "The second conjunct of the goal is expected to be a equality whose"
+ "right-hand side is the application of a function to a tuple of terms of the shape" shape
+ "but the right-hand side is not a function application:" RHS
+ end.
+Ltac check_precondition _ :=
+ lazymatch goal with
+ | [ |- ?is_bounded_by /\ ?LHS = ?RHS ]
+ => check_is_bounded_by_shape is_bounded_by;
+ check_LHS_Z_shape LHS;
+ check_RHS_Z_shape RHS
+ | [ |- ?G ]
+ => let shape := uconstr:(?is_bounded /\ ?LHS = ?RHS) in
+ fail "The goal has the wrong shape for reflective gluing; expected" shape "but found" G
+ end.
+Ltac split_BoundedWordToZ :=
+ (** first revert the context definition which is an evar named [f]
+ in the docs above, so that it becomes evar 1 (for
+ [instantiate]), and so that [make_evar_for_first_projection]
+ works. It's not the most robust way to find the right term;
+ maybe we should modify some of the checks above to assert that
+ the evar found is a particular one? *)
+ check_precondition ();
+ lazymatch goal with
+ | [ |- _ /\ ?LHS = _ ]
+ => match goal with
+ | [ f := ?e |- _ ]
+ => is_evar e; match LHS with context[f] => idtac end;
+ revert f
+ end
+ end;
+ repeat match goal with
+ | [ |- context[map wordToZ (proj1_sig ?x)] ]
+ => is_var x;
+ first [ clearbody x; fail 1
+ | (** we want to keep the same context variable in
+ the evar that we reverted above, and in the
+ current goal; hence the instantiate trick *)
+ instantiate (1:=ltac:(destruct x)); destruct x ]
+ | [ H := context[map wordToZ (proj1_sig ?x)] |- _ ]
+ => is_var x;
+ first [ clearbody x; fail 1
+ | (** we want to keep the same context variable in
+ the evar that we reverted above, and in the
+ current goal; hence the instantiate trick *)
+ instantiate (1:=ltac:(destruct x)); destruct x ]
+ | [ |- context[fst ?x] ]
+ => is_var x;
+ first [ clearbody x; fail 1
+ | (** we want to keep the same context variable in
+ the evar that we reverted above, and in the
+ current goal; hence the instantiate trick *)
+ change (fst x) with (let (a, b) := x in a);
+ change (snd x) with (let (a, b) := x in b);
+ instantiate (1:=ltac:(destruct x)); destruct x ];
+ cbv beta iota
+ end;
+ cbv beta iota in *; intro; (* put [f] back in the context so that [cbn] doesn't remove this let-in *)
+ cbn [proj1_sig] in *.
+
+(** ** [zrange_to_reflective] *)
+(** The [zrange_to_reflective] tactic takes a goal of the form
+<<
+(is_bounded_by _ bounds (map wordToZ (?fW args)) /\ ...)
+ /\ (map wordToZ (?fW args), ...) = fZ argsZ
+>>
+ and uses [cut] and a small lemma to turn it into a goal that the
+ reflective machinery can handle. The goal left by this tactic
+ should be fully solvable by the reflective pipeline. *)
+
+Lemma adjust_goal_for_reflective {T P} (LHS RHS : T)
+ : P RHS /\ LHS = RHS -> P LHS /\ LHS = RHS.
+Proof. intros [? ?]; subst; tauto. Qed.
+Ltac adjust_goal_for_reflective := apply adjust_goal_for_reflective.
+Ltac unmap_wordToZ_tuple term :=
+ lazymatch term with
+ | (?x, ?y) => let x' := unmap_wordToZ_tuple x in
+ let y' := unmap_wordToZ_tuple y in
+ constr:((x', y'))
+ | map wordToZ ?x => x
+ end.
+Ltac bounds_from_is_bounded_by T :=
+ lazymatch T with
+ | ?A /\ ?B => let a := bounds_from_is_bounded_by A in
+ let b := bounds_from_is_bounded_by B in
+ constr:((a, b))
+ | ZRange.is_bounded_by _ ?bounds _
+ => bounds
+ end.
+Ltac pose_proof_bounded_from_Zargs_hyps Zargs H :=
+ lazymatch Zargs with
+ | (?a, ?b)
+ => let Ha := fresh in
+ let Hb := fresh in
+ pose_proof_bounded_from_Zargs_hyps a Ha;
+ pose_proof_bounded_from_Zargs_hyps b Hb;
+ let pf := constr:(conj Ha Hb) in
+ lazymatch type of pf with
+ | @Bounds.is_bounded_by ?A ?boundsA (@cast_back_flat_const ?var ?tA ?f ?VA ?argsA)
+ /\ @Bounds.is_bounded_by ?B ?boundsB (@cast_back_flat_const ?var ?tB ?f ?VB ?argsB)
+ => pose proof
+ ((pf : @Bounds.is_bounded_by
+ (Prod A B) (boundsA, boundsB)
+ (@cast_back_flat_const var (Prod tA tB) f (VA, VB) (argsA, argsB))))
+ as H;
+ clear Ha Hb
+ | ?pfT
+ => let shape
+ := uconstr:(@Bounds.is_bounded_by ?A ?boundsA (@cast_back_flat_const ?var ?tA ?f ?VA ?argsA)
+ /\ @Bounds.is_bounded_by ?B ?boundsB (@cast_back_flat_const ?var ?tB ?f ?VB ?argsB)) in
+ fail 1 "Returned value from recursive call of bounded_from_Zargs_hyps has the wrong type"
+ "Cannot match type" pfT
+ "with shape" shape
+ end
+ | Tuple.map wordToZ ?arg
+ => lazymatch goal with
+ | [ H' : Bounds.is_bounded_by ?bounds (cast_back_flat_const arg) |- _ ]
+ => rename H' into H
+ | _ => let shape := uconstr:(Bounds.is_bounded_by _ (cast_back_flat_const arg)) in
+ idtac "In the context:"; print_context ();
+ fail 1 "Could not find bounds in the context for" arg
+ "when looking for a hypothesis of shape" shape
+ end
+ end.
+Ltac find_reified_f_evar LHS :=
+ lazymatch LHS with
+ | fst ?x => find_reified_f_evar x
+ | snd ?x => find_reified_f_evar x
+ | (?x, _) => find_reified_f_evar x
+ | map wordToZ ?x => find_reified_f_evar x
+ | _ => LHS
+ end.
+Ltac zrange_to_reflective_hyps_step :=
+ match goal with
+ | [ H : @ZRange.is_bounded_by ?option_bit_width ?count ?bounds (Tuple.map wordToZ ?arg) |- _ ]
+ => let rT := constr:(Syntax.tuple (Tbase TZ) count) in
+ let is_bounded_by' := constr:(@Bounds.is_bounded_by rT) in
+ let map' := constr:(@cast_back_flat_const (@Bounds.interp_base_type) rT (fun _ => Bounds.bounds_to_base_type) bounds) in
+ (* we use [assert] and [abstract] rather than [change] to catch
+ inefficiencies in conversion early, rather than allowing
+ [Defined] to take forever *)
+ let H' := fresh H in
+ rename H into H';
+ assert (H : is_bounded_by' bounds (map' arg)) by (clear -H'; abstract exact H');
+ clear H'; move H at top
+ | [ H := context Hv[@Tuple.map ?a ?b ?c (@wordToZ ?d) ?x], Hbounded : Bounds.is_bounded_by ?bounds (cast_back_flat_const ?x) |- _ ]
+ => let T := type of (@Tuple.map a b c (@wordToZ d) x) in
+ let T := (eval compute in T) in
+ let rT := reify_flat_type T in
+ let map_t := constr:(fun t bs => @cast_back_flat_const (@Bounds.interp_base_type) t (fun _ => Bounds.bounds_to_base_type) bs) in
+ let map' := constr:(map_t rT bounds) in
+ let Hv' := context Hv[map' x] in
+ progress change Hv' in (value of H); cbv beta in H
+ end.
+Ltac zrange_to_reflective_hyps := repeat zrange_to_reflective_hyps_step.
+Ltac zrange_to_reflective_goal Hbounded :=
+ lazymatch goal with
+ | [ |- ?is_bounded_by_T /\ ?LHS = ?f ?Zargs ]
+ => let T := type of f in
+ let f_domain := lazymatch eval hnf in T with ?A -> ?B => A end in
+ let T := (eval compute in T) in
+ let rT := reify_type T in
+ let is_bounded_by' := constr:(@Bounds.is_bounded_by (codomain rT)) in
+ let output_bounds := bounds_from_is_bounded_by is_bounded_by_T in
+ pose_proof_bounded_from_Zargs_hyps Zargs Hbounded;
+ let input_bounds := lazymatch type of Hbounded with Bounds.is_bounded_by ?bounds _ => bounds end in
+ let map_t := constr:(fun t bs => @cast_back_flat_const (@Bounds.interp_base_type) t (fun _ => Bounds.bounds_to_base_type) bs) in
+ let map_output := constr:(map_t (codomain rT) output_bounds) in
+ let map_input := constr:(map_t (domain rT) input_bounds) in
+ let args := unmap_wordToZ_tuple Zargs in
+ let reified_f_evar := find_reified_f_evar LHS in
+ (* we use [cut] and [abstract] rather than [change] to catch
+ inefficiencies in conversion early, rather than allowing
+ [Defined] to take forever *)
+ cut (is_bounded_by' output_bounds (map_output reified_f_evar) /\ map_output reified_f_evar = f (map_input args));
+ [ generalize reified_f_evar; clear; clearbody f; clear; let x := fresh in intros ? x; abstract exact x
+ | ];
+ cbv beta
+ end;
+ adjust_goal_for_reflective.
+Ltac zrange_to_reflective Hbounded := zrange_to_reflective_hyps; zrange_to_reflective_goal Hbounded.
+
+(** ** [refine_to_reflective_glue] *)
+(** The tactic [refine_to_reflective_glue] is the public-facing one;
+ it takes a goal of the form
+<<
+BoundedWordToZ ?f = F (BoundedWordToZ A) (BoundedWordToZ B) ... (BoundedWordToZ Z)
+>>
+ where [?f] is an evar, and turns it into a goal the that
+ reflective automation pipeline can handle. *)
+Ltac refine_to_reflective_glue' Hbounded :=
+ reassoc_sig_and_eexists;
+ do_curry_rhs;
+ split_BoundedWordToZ;
+ zrange_to_reflective Hbounded.
+Ltac refine_to_reflective_glue :=
+ let Hbounded := fresh "Hbounded" in
+ refine_to_reflective_glue' Hbounded.
diff --git a/src/Compilers/Z/Bounds/Pipeline/OutputType.v b/src/Compilers/Z/Bounds/Pipeline/OutputType.v
new file mode 100644
index 000000000..8205ef70c
--- /dev/null
+++ b/src/Compilers/Z/Bounds/Pipeline/OutputType.v
@@ -0,0 +1,51 @@
+(** * Definition of the output type of the post-Wf pipeline *)
+(** Do not change these definitions unless you're hacking on the
+ entire reflective pipeline tactic automation. *)
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.SmartMap.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Bounds.Interpretation.
+Require Import Crypto.Util.Sigma.
+Require Import Crypto.Util.Prod.
+Local Notation pick_typeb := Bounds.bounds_to_base_type (only parsing).
+Local Notation pick_type v := (SmartFlatTypeMap (fun _ => pick_typeb) v).
+
+Record ProcessedReflectivePackage
+ := { InputType : _;
+ input_expr : Expr base_type op InputType;
+ input_bounds : interp_flat_type Bounds.interp_base_type (domain InputType);
+ output_bounds :> interp_flat_type Bounds.interp_base_type (codomain InputType);
+ output_expr :> Expr base_type op (Arrow (pick_type input_bounds) (pick_type output_bounds)) }.
+
+Notation OutputType pkg
+ := (Arrow (pick_type (@input_bounds pkg)) (pick_type (@output_bounds pkg)))
+ (only parsing).
+
+Definition Build_ProcessedReflectivePackage_from_option_sigma
+ {t} (e : Expr base_type op t)
+ (input_bounds : interp_flat_type Bounds.interp_base_type (domain t))
+ (result : option { output_bounds : interp_flat_type Bounds.interp_base_type (codomain t)
+ & Expr base_type op (Arrow (pick_type input_bounds) (pick_type output_bounds)) })
+ : option ProcessedReflectivePackage
+ := option_map
+ (fun be
+ => let 'existT b e' := be in
+ {| InputType := t ; input_expr := e ; input_bounds := input_bounds ; output_bounds := b ; output_expr := e' |})
+ result.
+
+Definition ProcessedReflectivePackage_to_sigT (x : ProcessedReflectivePackage)
+ : { InputType : _
+ & Expr base_type op InputType
+ * { bounds : interp_flat_type Bounds.interp_base_type (domain InputType)
+ * interp_flat_type Bounds.interp_base_type (codomain InputType)
+ & Expr base_type op (Arrow (pick_type (fst bounds)) (pick_type (snd bounds))) } }%type
+ := let (a, b, c, d, e) := x in
+ existT _ a (b, (existT _ (c, d) e)).
+
+Ltac inversion_ProcessedReflectivePackage :=
+ repeat match goal with
+ | [ H : _ = _ :> ProcessedReflectivePackage |- _ ]
+ => apply (f_equal ProcessedReflectivePackage_to_sigT) in H;
+ cbv [ProcessedReflectivePackage_to_sigT] in H
+ end;
+ inversion_sigma; inversion_prod.
diff --git a/src/Compilers/Z/Bounds/Pipeline/ReflectiveTactics.v b/src/Compilers/Z/Bounds/Pipeline/ReflectiveTactics.v
new file mode 100644
index 000000000..3586cc78d
--- /dev/null
+++ b/src/Compilers/Z/Bounds/Pipeline/ReflectiveTactics.v
@@ -0,0 +1,288 @@
+(** * Reflective Pipeline: Tactics that execute the pipeline *)
+(** N.B. This file should not need to be changed in normal
+ modifications of the reflective transformations; to modify the
+ transformations performed in the reflective pipeline; see
+ Pipeline/Definition.v. If the input format of the pre-reflective
+ goal changes, prefer adding complexity to Pipeline/Glue.v to
+ transform the goal and hypotheses into a uniform syntax to
+ modifying this file. This file will need to be modified if you
+ perform heavy changes in the shape of the generic or ℤ-specific
+ reflective machinery itself, or if you find bugs or slowness. *)
+(** ** Preamble *)
+Require Import Coq.ZArith.ZArith.
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Wf.
+Require Import Crypto.Compilers.WfReflective.
+Require Import Crypto.Compilers.RenameBinders.
+Require Import Crypto.Compilers.Eta.
+Require Import Crypto.Compilers.EtaInterp.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.Util.
+Require Import Crypto.Compilers.Z.Bounds.Interpretation.
+Require Import Crypto.Compilers.Z.Bounds.Relax.
+Require Import Crypto.Compilers.Reify.
+Require Import Crypto.Compilers.Z.Reify.
+Require Import Crypto.Compilers.Z.Bounds.Pipeline.Definition.
+Require Import Crypto.Util.Tactics.Head.
+Require Import Crypto.Util.Tactics.SubstLet.
+Require Import Crypto.Util.Tactics.UnifyAbstractReflexivity.
+Require Import Crypto.Util.FixedWordSizes.
+Require Import Crypto.Util.Option.
+Require Import Bedrock.Word.
+
+(** The final tactic in this file, [do_reflective_pipeline], takes a
+ goal of the form
+<<
+@Bounds.is_bounded_by (codomain T) bounds (fZ (cast_back_flat_const v))
+ /\ cast_back_flat_const fW = fZ (cast_back_flat_const v)
+>>
+
+ where [fW] must be a context definition which is a single evar,
+ and all other terms must be evar-free. It fully solves the goal,
+ instantiating [fW] with an appropriately-unfolded
+ (reflection-definition-free) version of [fZ (cast_back_flat_const
+ v)] which has been transformed by the reflective pipeline. *)
+
+Module Export Exports.
+ Export Crypto.Compilers.Reify. (* export for the instances for recursing under binders *)
+ Export Crypto.Compilers.Z.Reify. (* export for the tactic redefinitions *)
+ Export Crypto.Compilers.Z.Bounds.Pipeline.Definition.Exports.
+End Exports.
+
+(** ** Reification *)
+(** The [do_reify] tactic handles goals of the form
+<<
+forall x, Interp _ ?e x = F
+>>
+ by reifying [F]. *)
+Ltac do_reify :=
+ cbv beta iota delta [Tuple.tuple Tuple.tuple'] in *;
+ cbv beta iota delta [Syntax.interp_flat_type Syntax.interp_base_type];
+ reify_context_variables;
+ Reify_rhs; reflexivity.
+(** ** Input Boundedness Side-Conditions *)
+(** The tactic [handle_bounds_from_hyps] handles goals of the form
+<<
+Bounds.is_bounded_by (_, _, ..., _) _
+>>
+ by splitting them apart and looking in the context for hypotheses
+ that prove the bounds. *)
+Ltac handle_bounds_from_hyps :=
+ repeat match goal with
+ | _ => assumption
+ | [ |- cast_back_flat_const _ = cast_back_flat_const _ ] => reflexivity
+ | [ |- _ /\ _ ] => split
+ | [ |- Bounds.is_bounded_by (_, _) _ ] => split
+ end.
+(** ** Unfolding [Interp] *)
+(** The reduction strategies [interp_red], [extra_interp_red], and
+ [constant_simplification] (the latter two defined in
+ Pipeline/Definition.v) define the constants that get unfolded
+ before instantiating the original evar with [Interp _
+ vm_computed_reified_expression arguments]. *)
+Declare Reduction interp_red
+ := cbv [fst snd
+ Interp InterpEta interp_op interp interp_eta interpf interpf_step
+ interp_flat_type_eta interp_flat_type_eta_gen interp_flat_type
+ interp_base_type interp_op
+ SmartMap.SmartFlatTypeMap SmartMap.SmartFlatTypeMapUnInterp SmartMap.SmartFlatTypeMapInterp2
+ SmartMap.smart_interp_flat_map
+ codomain domain
+ lift_op Zinterp_op cast_const
+ ZToInterp interpToZ
+ ].
+
+(** ** Solving Side-Conditions of Equality *)
+(** This section defines a number of different ways to solve goals of
+ the form [LHS = RHS] where [LHS] may contain evars and [RHS] must
+ not contain evars. Most tactics use [abstract] to reduce the load
+ on [Defined] and to catch looping behavior early. *)
+
+(** The tactic [unify_abstract_renamify_rhs_reflexivity] calls [renamify] on [RHS] and unifies
+ that with [LHS]; and then costs one [vm_compute] to prove the
+ equality. *)
+Ltac unify_abstract_renamify_rhs_reflexivity :=
+ unify_transformed_rhs_abstract_tac
+ ltac:(renamify)
+ unify_tac
+ vm_cast_no_check.
+(** The tactic [unify_abstract_cbv_interp_rhs_reflexivity] runs the interpretation
+ reduction strategies in [RHS] and unifies the result with [LHS],
+ and does not use the vm (and hence does not fully reduce things,
+ which is important for efficiency). *)
+Ltac unify_abstract_cbv_interp_rhs_reflexivity :=
+ intros; clear;
+ lazymatch goal with
+ | [ |- ?LHS = ?RHS ]
+ => let RHS' := (eval interp_red in RHS) in
+ let RHS' := (eval extra_interp_red in RHS') in
+ let RHS' := lazymatch do_constant_simplification with
+ | true => (eval constant_simplification in RHS')
+ | _ => RHS'
+ end in
+ unify LHS RHS'; abstract exact_no_check (eq_refl RHS')
+ end.
+
+
+(** ** Assemble the parts of Pipeline.Definition, in Gallina *)
+(** In this section, we assemble [PreWfPipeline] and [PostWfPipeline],
+ and add extra equality hypotheses to minimize the work we have to
+ do in Ltac. *)
+(** *** Gallina assembly imports *)
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.SmartMap.
+Require Import Crypto.Compilers.Wf.
+Require Import Crypto.Compilers.WfReflectiveGen.
+Require Import Crypto.Compilers.WfReflective.
+Require Import Crypto.Compilers.Eta.
+Require Import Crypto.Compilers.EtaWf.
+Require Import Crypto.Compilers.EtaInterp.
+Require Import Crypto.Compilers.Z.Bounds.Pipeline.OutputType.
+Require Import Crypto.Compilers.Z.Bounds.Pipeline.Definition.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.Equality.
+Require Import Crypto.Compilers.Z.Syntax.Util.
+Require Import Crypto.Compilers.Z.Bounds.Interpretation.
+Require Import Crypto.Compilers.Z.Bounds.Relax.
+Require Import Crypto.Util.PartiallyReifiedProp.
+Require Import Crypto.Util.Equality.
+
+(** *** Gallina assembly *)
+Local Notation pick_typeb := Bounds.bounds_to_base_type (only parsing).
+Local Notation pick_type v := (SmartFlatTypeMap (fun _ => pick_typeb) v).
+Definition PipelineCorrect
+ {t}
+ {input_bounds : interp_flat_type Bounds.interp_base_type (domain t)}
+ {given_output_bounds : interp_flat_type Bounds.interp_base_type (codomain t)}
+ {v' : interp_flat_type Syntax.interp_base_type (pick_type input_bounds)}
+ {b e' e_final e_final_newtype}
+ {fZ}
+ {final_e_evar : interp_flat_type Syntax.interp_base_type (pick_type given_output_bounds)}
+ {e}
+ {e_pkg}
+ (** ** reification *)
+ (rexpr_sig : { rexpr : Expr base_type op t | forall x, Interp Syntax.interp_op rexpr x = fZ x })
+ (** ** pre-wf pipeline *)
+ (He : e = PreWfPipeline (proj1_sig rexpr_sig))
+ (** ** proving wf *)
+ (He_unnatize_for_wf : forall var, unnatize_expr 0 (ExprEta' e (fun t => (nat * var t)%type)) = ExprEta' e _)
+ (Hwf : forall var1 var2,
+ let P := (@reflect_wfT base_type base_type_eq_semidec_transparent op op_beq var1 var2 nil _ _ (ExprEta' e _) (ExprEta' e _)) in
+ trueify P = P)
+ (** ** post-wf-pipeline *)
+ (Hpost : e_pkg = PostWfPipeline e input_bounds)
+ (Hpost_correct : Some {| input_expr := e ; input_bounds := input_bounds ; output_bounds := b ; output_expr := e' |} = e_pkg)
+ (** ** renaming *)
+ (Hrenaming : e_final = e')
+ (** ** bounds relaxation *)
+ (Hbounds_sane : pick_type given_output_bounds = pick_type b)
+ (Hbounds_relax : Bounds.is_tighter_thanb b given_output_bounds = true)
+ (Hbounds_sane_refl
+ : e_final_newtype
+ = eq_rect _ (fun t => Expr base_type op (Arrow (pick_type input_bounds) t)) e' _ (eq_sym Hbounds_sane))
+ (** ** instantiation of original evar *)
+ (Hevar : final_e_evar = Interp (t:=Arrow _ _) Syntax.interp_op e_final_newtype v')
+ (** ** side condition *)
+ (Hv : Bounds.is_bounded_by input_bounds (cast_back_flat_const v'))
+ : Bounds.is_bounded_by given_output_bounds (fZ (cast_back_flat_const v'))
+ /\ cast_back_flat_const final_e_evar = fZ (cast_back_flat_const v').
+Proof.
+ destruct rexpr_sig as [? Hrexpr].
+ assert (Hwf' : Wf e)
+ by (apply (proj1 (@Wf_ExprEta'_iff _ _ _ e));
+ eapply reflect_Wf;
+ [ .. | intros; split; [ eapply He_unnatize_for_wf | rewrite <- Hwf; apply trueify_true ] ];
+ auto using base_type_eq_semidec_is_dec, op_beq_bl).
+ clear Hwf He_unnatize_for_wf.
+ symmetry in Hpost_correct.
+ subst; cbv [proj1_sig] in *.
+ rewrite <- Hrexpr.
+ eapply PostWfPipelineCorrect in Hpost_correct; [ | solve [ eauto ].. ].
+ rewrite !@InterpPreWfPipeline in Hpost_correct.
+ unshelve eapply relax_output_bounds; try eassumption; [].
+ match goal with
+ | [ |- context[Interp _ (@eq_rect ?A ?x ?P ?k ?y ?pf) ?v] ]
+ => rewrite (@ap_transport A P _ x y pf (fun t e => Interp interp_op e v) k)
+ end.
+ rewrite <- transport_pp, concat_Vp; simpl.
+ apply Hpost_correct.
+Qed.
+
+
+(** ** Assembling the Pipeline, in Ltac *)
+(** The tactic [refine_with_pipeline_correct] uses the
+ [PipelineCorrect] lemma to create side-conditions. It assumes the
+ goal is in exactly the form given in the conclusion of the
+ [PipelineCorrect] lemma. *)
+Ltac refine_with_pipeline_correct :=
+ lazymatch goal with
+ | [ |- _ /\ ?castback ?fW = ?fZ ?arg ]
+ => let lem := open_constr:(@PipelineCorrect _ _ _ _ _ _ _ _ _ _ _ _) in
+ simple refine (lem _ _ _ _ _ _ _ _ _ _ _ _);
+ subst fW fZ
+ end;
+ [ eexists
+ | cbv [proj1_sig].. ].
+
+(** The tactic [solve_side_conditions] uses the above
+ reduction-and-proving-equality tactics to prove the
+ side-conditions of [PipelineCorrect]. The order must match with
+ [PipelineCorrect]. Which tactic to use was chosen in the
+ following way:
+
+ - The default is [unify_abstract_vm_compute_rhs_reflexivity]
+
+ - If the [RHS] is already in [vm_compute]d form, use
+ [unify_abstract_rhs_reflexivity] (saves a needless [vm_compute] which would be a
+ costly no-op)
+
+ - If the proof needs to be transparent and there are no evars and
+ you want the user to see the fully [vm_compute]d term on error,
+ use [vm_compute; reflexivity]
+
+ - If the user should see an unreduced term and you're proving [_ =
+ true], use [abstract vm_cast_no_check (eq_refl true)]
+
+ - If you want to preserve binder names, use [unify_abstract_cbv_rhs_reflexivity]
+
+ The other choices are tactics that are specialized to the specific
+ side-condition for which they are used (reification, boundedness
+ of input, reduction of [Interp], renaming). *)
+Ltac solve_side_conditions :=
+ [>
+ (** ** reification *)
+ do_reify |
+ (** ** pre-wf pipeline *)
+ unify_abstract_vm_compute_rhs_reflexivity |
+ (** ** reflective wf side-condition 1 *)
+ unify_abstract_vm_compute_rhs_reflexivity |
+ (** ** reflective wf side-condition 2 *)
+ unify_abstract_vm_compute_rhs_reflexivity |
+ (** ** post-wf pipeline *)
+ unify_abstract_vm_compute_rhs_reflexivity |
+ (** ** post-wf pipeline gives [Some _] *)
+ unify_abstract_rhs_reflexivity |
+ (** ** renaming binders *)
+ unify_abstract_renamify_rhs_reflexivity |
+ (** ** types computed from given output bounds are the same as types computed from computed output bounds *)
+ (** N.B. the proof must be exactly [eq_refl] because it's used in a
+ later goal and needs to reduce *)
+ subst_let; clear; vm_compute; reflexivity |
+ (** ** computed output bounds are not looser than the given output bounds *)
+ (** we do subst and we don't [vm_compute] first because we want to
+ get an error message that displays the bounds *)
+ subst_let; clear; abstract vm_cast_no_check (eq_refl true) |
+ (** ** removal of a cast across the equality proof above *)
+ unify_abstract_compute_rhs_reflexivity |
+ (** ** unfolding of [interp] constants *)
+ unify_abstract_cbv_interp_rhs_reflexivity |
+ (** ** boundedness of inputs *)
+ abstract handle_bounds_from_hyps ].
+
+
+(** ** The Entire Pipeline *)
+(** The [do_reflective_pipeline] tactic solves a goal of the form that
+ is described at the top of this file, and is the public interface
+ of this file. *)
+Ltac do_reflective_pipeline :=
+ refine_with_pipeline_correct; solve_side_conditions.
diff --git a/src/Compilers/Z/Bounds/Relax.v b/src/Compilers/Z/Bounds/Relax.v
new file mode 100644
index 000000000..40b678071
--- /dev/null
+++ b/src/Compilers/Z/Bounds/Relax.v
@@ -0,0 +1,127 @@
+Require Import Coq.ZArith.ZArith.
+Require Import Coq.Classes.Morphisms.
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.TypeInversion.
+Require Import Crypto.Compilers.Relations.
+Require Import Crypto.Compilers.SmartMap.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.Equality.
+Require Import Crypto.Compilers.Z.Syntax.Util.
+Require Import Crypto.Compilers.Z.Bounds.Interpretation.
+Require Import Crypto.Util.Tactics.DestructHead.
+Require Import Crypto.Util.Tactics.SpecializeBy.
+Require Import Crypto.Util.Tactics.BreakMatch.
+Require Import Crypto.Util.Tactics.SplitInContext.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Bool.
+
+Local Lemma helper logsz v
+ : (v < 2 ^ 2 ^ Z.of_nat logsz)%Z <-> (Z.to_nat (Z.log2_up (Z.log2_up (1 + v))) <= logsz)%nat.
+Proof.
+ rewrite Nat2Z.inj_le, Z2Nat.id by auto with zarith.
+ transitivity (1 + v <= 2^2^Z.of_nat logsz)%Z; [ omega | ].
+ rewrite !Z.log2_up_le_pow2_full by auto with zarith.
+ reflexivity.
+Qed.
+
+Local Arguments Z.pow : simpl never.
+Local Arguments Z.sub !_ !_.
+Local Arguments Z.add !_ !_.
+Local Arguments Z.mul !_ !_.
+Lemma relax_output_bounds'
+ t (tight_output_bounds relaxed_output_bounds : interp_flat_type Bounds.interp_base_type t)
+ (Hv : SmartFlatTypeMap (fun _ => Bounds.bounds_to_base_type) relaxed_output_bounds
+ = SmartFlatTypeMap (fun _ => Bounds.bounds_to_base_type) tight_output_bounds)
+ v k
+ (v' := eq_rect _ (interp_flat_type _) v _ Hv)
+ (Htighter : @Bounds.is_bounded_by
+ t tight_output_bounds
+ (@cast_back_flat_const
+ (@Bounds.interp_base_type) t (fun _ => Bounds.bounds_to_base_type) tight_output_bounds
+ v')
+ /\ @cast_back_flat_const
+ (@Bounds.interp_base_type) t (fun _ => Bounds.bounds_to_base_type) tight_output_bounds
+ v'
+ = k)
+ (Hrelax : Bounds.is_tighter_thanb tight_output_bounds relaxed_output_bounds = true)
+ : @Bounds.is_bounded_by
+ t relaxed_output_bounds
+ (@cast_back_flat_const
+ (@Bounds.interp_base_type) t (fun _ => Bounds.bounds_to_base_type) relaxed_output_bounds
+ v)
+ /\ @cast_back_flat_const
+ (@Bounds.interp_base_type) t (fun _ => Bounds.bounds_to_base_type) relaxed_output_bounds
+ v
+ = k.
+Proof.
+ destruct Htighter as [H0 H1]; subst v' k.
+ cbv [Bounds.is_bounded_by cast_back_flat_const Bounds.is_tighter_thanb] in *.
+ apply interp_flat_type_rel_pointwise_iff_relb in Hrelax.
+ induction t; unfold SmartFlatTypeMap in *; simpl @smart_interp_flat_map in *; inversion_flat_type.
+ { cbv [Bounds.is_tighter_thanb' ZRange.is_tighter_than_bool is_true SmartFlatTypeMap Bounds.bounds_to_base_type ZRange.is_bounded_by' ZRange.is_bounded_by Bounds.is_bounded_by' Bounds.bit_width_of_base_type] in *; simpl in *.
+ repeat first [ progress inversion_flat_type
+ | progress inversion_base_type
+ | progress destruct_head bounds
+ | progress split_andb
+ | progress Z.ltb_to_lt
+ | progress break_match_hyps
+ | progress destruct_head'_and
+ | progress simpl in *
+ | rewrite helper in *
+ | omega
+ | tauto
+ | congruence
+ | progress destruct_head @eq; (reflexivity || omega)
+ | progress break_innermost_match_step
+ | apply conj ]. }
+ { compute in *; tauto. }
+ { simpl in *.
+ specialize (fun Hv => IHt1 (fst tight_output_bounds) (fst relaxed_output_bounds) Hv (fst v)).
+ specialize (fun Hv => IHt2 (snd tight_output_bounds) (snd relaxed_output_bounds) Hv (snd v)).
+ do 2 match goal with
+ | [ H : _ = _, H' : forall x, _ |- _ ] => specialize (H' H)
+ end.
+ simpl in *.
+ split_and.
+ repeat apply conj;
+ [ match goal with H : _ |- _ => apply H end..
+ | apply (f_equal2 (@pair _ _)); (etransitivity; [ match goal with H : _ |- _ => apply H end | ]) ];
+ repeat first [ progress destruct_head prod
+ | progress simpl in *
+ | reflexivity
+ | assumption
+ | match goal with
+ | [ |- ?P (eq_rect _ _ _ _ _) = ?P _ ]
+ => apply f_equal; clear
+ | [ H : interp_flat_type_rel_pointwise (@Bounds.is_bounded_by') ?x ?y |- interp_flat_type_rel_pointwise (@Bounds.is_bounded_by') ?x ?y' ]
+ => clear -H;
+ match goal with |- ?R _ _ => generalize dependent R; intros end
+ | [ H : ?x = ?y |- _ ]
+ => first [ generalize dependent x | generalize dependent y ];
+ let k := fresh in intro k; intros; subst k
+ end ]. }
+Qed.
+
+Lemma relax_output_bounds
+ t (tight_output_bounds relaxed_output_bounds : interp_flat_type Bounds.interp_base_type t)
+ (Hv : SmartFlatTypeMap (fun _ => Bounds.bounds_to_base_type) relaxed_output_bounds
+ = SmartFlatTypeMap (fun _ => Bounds.bounds_to_base_type) tight_output_bounds)
+ v k
+ (v' := eq_rect _ (interp_flat_type _) v _ Hv)
+ (Htighter : @Bounds.is_bounded_by t tight_output_bounds k
+ /\ @cast_back_flat_const
+ (@Bounds.interp_base_type) t (fun _ => Bounds.bounds_to_base_type) tight_output_bounds
+ v'
+ = k)
+ (Hrelax : Bounds.is_tighter_thanb tight_output_bounds relaxed_output_bounds = true)
+ : @Bounds.is_bounded_by t relaxed_output_bounds k
+ /\ @cast_back_flat_const
+ (@Bounds.interp_base_type) t (fun _ => Bounds.bounds_to_base_type) relaxed_output_bounds
+ v
+ = k.
+Proof.
+ pose proof (fun pf => @relax_output_bounds' t tight_output_bounds relaxed_output_bounds Hv v k (conj pf (proj2 Htighter)) Hrelax) as H.
+ destruct H as [H1 H2]; [ | rewrite <- H2; tauto ].
+ subst v'.
+ destruct Htighter; subst k; assumption.
+Qed.
diff --git a/src/Compilers/Z/CNotations.v b/src/Compilers/Z/CNotations.v
new file mode 100644
index 000000000..8ec885db3
--- /dev/null
+++ b/src/Compilers/Z/CNotations.v
@@ -0,0 +1,773 @@
+Require Export Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Export Crypto.Compilers.Z.HexNotationConstants.
+Require Export Crypto.Util.Notations.
+
+Reserved Notation "T x = A ; b" (at level 200, b at level 200, format "T x = A ; '//' b").
+Reserved Notation "x & y" (at level 40).
+
+Global Open Scope expr_scope.
+
+Notation "T x = A ; b" := (LetIn (tx:=T) A (fun x => b)) : expr_scope.
+(* python:
+<<
+types = ('bool', 'uint8_t', 'uint8_t', 'uint8_t', 'uint16_t', 'uint32_t', 'uint64_t', 'uint128_t')
+for lgwordsz in range(0, len(types)):
+ print('Notation "\'%s\'" := (Tbase (TWord %d)).' % (types[lgwordsz], lgwordsz))
+print('Notation ℤ := (Tbase TZ).')
+print('')
+cast_pat = "'(%s)' %s"
+for opn, op, lvl in (('*', 'Mul', 40), ('+', 'Add', 50), ('-', 'Sub', 50), ('&', 'Land', 40), ('<<', 'Shl', 30)):
+ for v1 in (False, True):
+ for v2 in (False, True):
+ lhs = ('x' if not v1 else '(Var x)')
+ rhs = ('y' if not v2 else '(Var y)')
+ print('Notation "x %s y" := (Op (%s _ _ _) (Pair %s %s)).' % (opn, op, lhs, rhs))
+ for lgwordsz in range(0, len(types)):
+ for v1 in (False, True):
+ for v2 in (False, True):
+ lhs = ('x' if not v1 else '(Var x)')
+ rhs = ('y' if not v2 else '(Var y)')
+ print('Notation "%s %s %s" := (Op (%s (TWord _) (TWord _) (TWord %d)) (Pair %s %s)) (at level %d).'
+ % (cast_pat % (types[lgwordsz], 'x'), opn, cast_pat % (types[lgwordsz], 'y'),
+ op, lgwordsz, lhs, rhs, lvl))
+ for v1 in (False, True):
+ for v2 in (False, True):
+ lhs = ('x' if not v1 else '(Var x)')
+ rhs = ('y' if not v2 else '(Var y)')
+ print('Notation "%s %s %s" := (Op (%s (TWord %d) (TWord _) (TWord %d)) (Pair %s %s)) (at level %d).'
+ % ('x', opn, cast_pat % (types[lgwordsz], 'y'),
+ op, lgwordsz, lgwordsz, lhs, rhs, lvl))
+ print('Notation "%s %s %s" := (Op (%s (TWord _) (TWord %d) (TWord %d)) (Pair %s %s)) (at level %d).'
+ % (cast_pat % (types[lgwordsz], 'x'), opn, 'y',
+ op, lgwordsz, lgwordsz, lhs, rhs, lvl))
+ for v1 in (False, True):
+ for v2 in (False, True):
+ lhs = ('x' if not v1 else '(Var x)')
+ rhs = ('y' if not v2 else '(Var y)')
+ print('Notation "x %s y" := (Op (%s (TWord %d) (TWord %d) (TWord %d)) (Pair %s %s)).'
+ % (opn, op, lgwordsz, lgwordsz, lgwordsz, lhs, rhs))
+for opn, op, lvl in (('>>', 'Shr', 30),):
+ for v1 in (False, True):
+ for v2 in (False, True):
+ lhs = ('x' if not v1 else '(Var x)')
+ rhs = ('y' if not v2 else '(Var y)')
+ print('Notation "x %s y" := (Op (%s _ _ _) (Pair %s %s)).' % (opn, op, lhs, rhs))
+ for lgwordsz in range(0, len(types)):
+ for v1 in (False, True):
+ for v2 in (False, True):
+ lhs = ('x' if not v1 else '(Var x)')
+ rhs = ('y' if not v2 else '(Var y)')
+ print('Notation "\'(%s)\' ( x %s y )" := (Op (%s (TWord _) (TWord _) (TWord %d)) (Pair %s %s)) (at level %d).'
+ % (types[lgwordsz], opn, op, lgwordsz, lhs, rhs, lvl))
+print('Notation Return x := (Var x).')
+print('Notation C_like := (Expr base_type op _).')
+>> *)
+Notation "'bool'" := (Tbase (TWord 0)).
+Notation "'uint8_t'" := (Tbase (TWord 1)).
+Notation "'uint8_t'" := (Tbase (TWord 2)).
+Notation "'uint8_t'" := (Tbase (TWord 3)).
+Notation "'uint16_t'" := (Tbase (TWord 4)).
+Notation "'uint32_t'" := (Tbase (TWord 5)).
+Notation "'uint64_t'" := (Tbase (TWord 6)).
+Notation "'uint128_t'" := (Tbase (TWord 7)).
+Notation ℤ := (Tbase TZ).
+
+Notation "x * y" := (Op (Mul _ _ _) (Pair x y)).
+Notation "x * y" := (Op (Mul _ _ _) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul _ _ _) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul _ _ _) (Pair (Var x) (Var y))).
+Notation "'(bool)' x * '(bool)' y" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 40).
+Notation "'(bool)' x * '(bool)' y" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40).
+Notation "'(bool)' x * '(bool)' y" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40).
+Notation "'(bool)' x * '(bool)' y" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(bool)' y" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 40).
+Notation "'(bool)' x * y" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 40).
+Notation "x * '(bool)' y" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40).
+Notation "'(bool)' x * y" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 40).
+Notation "x * '(bool)' y" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40).
+Notation "'(bool)' x * y" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(bool)' y" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(bool)' x * y" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 0) (TWord 0) (TWord 0)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x * '(uint8_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 40).
+Notation "'(uint8_t)' x * '(uint8_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40).
+Notation "'(uint8_t)' x * '(uint8_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint8_t)' x * '(uint8_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(uint8_t)' y" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 40).
+Notation "'(uint8_t)' x * y" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 40).
+Notation "x * '(uint8_t)' y" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40).
+Notation "'(uint8_t)' x * y" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 40).
+Notation "x * '(uint8_t)' y" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint8_t)' x * y" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(uint8_t)' y" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint8_t)' x * y" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 1) (TWord 1) (TWord 1)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x * '(uint8_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 40).
+Notation "'(uint8_t)' x * '(uint8_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40).
+Notation "'(uint8_t)' x * '(uint8_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint8_t)' x * '(uint8_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(uint8_t)' y" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 40).
+Notation "'(uint8_t)' x * y" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 40).
+Notation "x * '(uint8_t)' y" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40).
+Notation "'(uint8_t)' x * y" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 40).
+Notation "x * '(uint8_t)' y" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint8_t)' x * y" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(uint8_t)' y" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint8_t)' x * y" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 2) (TWord 2) (TWord 2)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x * '(uint8_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 40).
+Notation "'(uint8_t)' x * '(uint8_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40).
+Notation "'(uint8_t)' x * '(uint8_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint8_t)' x * '(uint8_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(uint8_t)' y" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 40).
+Notation "'(uint8_t)' x * y" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 40).
+Notation "x * '(uint8_t)' y" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40).
+Notation "'(uint8_t)' x * y" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 40).
+Notation "x * '(uint8_t)' y" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint8_t)' x * y" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(uint8_t)' y" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint8_t)' x * y" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 3) (TWord 3) (TWord 3)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))).
+Notation "'(uint16_t)' x * '(uint16_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 40).
+Notation "'(uint16_t)' x * '(uint16_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40).
+Notation "'(uint16_t)' x * '(uint16_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint16_t)' x * '(uint16_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(uint16_t)' y" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 40).
+Notation "'(uint16_t)' x * y" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 40).
+Notation "x * '(uint16_t)' y" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40).
+Notation "'(uint16_t)' x * y" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 40).
+Notation "x * '(uint16_t)' y" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint16_t)' x * y" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(uint16_t)' y" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint16_t)' x * y" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 4) (TWord 4) (TWord 4)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))).
+Notation "'(uint32_t)' x * '(uint32_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 40).
+Notation "'(uint32_t)' x * '(uint32_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40).
+Notation "'(uint32_t)' x * '(uint32_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint32_t)' x * '(uint32_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(uint32_t)' y" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 40).
+Notation "'(uint32_t)' x * y" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 40).
+Notation "x * '(uint32_t)' y" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40).
+Notation "'(uint32_t)' x * y" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 40).
+Notation "x * '(uint32_t)' y" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint32_t)' x * y" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(uint32_t)' y" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint32_t)' x * y" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 5) (TWord 5) (TWord 5)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))).
+Notation "'(uint64_t)' x * '(uint64_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 40).
+Notation "'(uint64_t)' x * '(uint64_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40).
+Notation "'(uint64_t)' x * '(uint64_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint64_t)' x * '(uint64_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(uint64_t)' y" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 40).
+Notation "'(uint64_t)' x * y" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 40).
+Notation "x * '(uint64_t)' y" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40).
+Notation "'(uint64_t)' x * y" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 40).
+Notation "x * '(uint64_t)' y" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint64_t)' x * y" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(uint64_t)' y" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint64_t)' x * y" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 6) (TWord 6) (TWord 6)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))).
+Notation "'(uint128_t)' x * '(uint128_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 40).
+Notation "'(uint128_t)' x * '(uint128_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40).
+Notation "'(uint128_t)' x * '(uint128_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint128_t)' x * '(uint128_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(uint128_t)' y" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 40).
+Notation "'(uint128_t)' x * y" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 40).
+Notation "x * '(uint128_t)' y" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40).
+Notation "'(uint128_t)' x * y" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 40).
+Notation "x * '(uint128_t)' y" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint128_t)' x * y" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(uint128_t)' y" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint128_t)' x * y" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 7) (TWord 7) (TWord 7)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))).
+Notation "x + y" := (Op (Add _ _ _) (Pair x y)).
+Notation "x + y" := (Op (Add _ _ _) (Pair x (Var y))).
+Notation "x + y" := (Op (Add _ _ _) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add _ _ _) (Pair (Var x) (Var y))).
+Notation "'(bool)' x + '(bool)' y" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 50).
+Notation "'(bool)' x + '(bool)' y" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50).
+Notation "'(bool)' x + '(bool)' y" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50).
+Notation "'(bool)' x + '(bool)' y" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(bool)' y" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 50).
+Notation "'(bool)' x + y" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 50).
+Notation "x + '(bool)' y" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50).
+Notation "'(bool)' x + y" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 50).
+Notation "x + '(bool)' y" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50).
+Notation "'(bool)' x + y" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(bool)' y" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(bool)' x + y" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 0) (TWord 0) (TWord 0)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x + '(uint8_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 50).
+Notation "'(uint8_t)' x + '(uint8_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50).
+Notation "'(uint8_t)' x + '(uint8_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint8_t)' x + '(uint8_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(uint8_t)' y" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 50).
+Notation "'(uint8_t)' x + y" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 50).
+Notation "x + '(uint8_t)' y" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50).
+Notation "'(uint8_t)' x + y" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 50).
+Notation "x + '(uint8_t)' y" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint8_t)' x + y" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(uint8_t)' y" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint8_t)' x + y" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 1) (TWord 1) (TWord 1)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x + '(uint8_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 50).
+Notation "'(uint8_t)' x + '(uint8_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50).
+Notation "'(uint8_t)' x + '(uint8_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint8_t)' x + '(uint8_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(uint8_t)' y" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 50).
+Notation "'(uint8_t)' x + y" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 50).
+Notation "x + '(uint8_t)' y" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50).
+Notation "'(uint8_t)' x + y" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 50).
+Notation "x + '(uint8_t)' y" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint8_t)' x + y" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(uint8_t)' y" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint8_t)' x + y" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 2) (TWord 2) (TWord 2)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x + '(uint8_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 50).
+Notation "'(uint8_t)' x + '(uint8_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50).
+Notation "'(uint8_t)' x + '(uint8_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint8_t)' x + '(uint8_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(uint8_t)' y" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 50).
+Notation "'(uint8_t)' x + y" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 50).
+Notation "x + '(uint8_t)' y" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50).
+Notation "'(uint8_t)' x + y" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 50).
+Notation "x + '(uint8_t)' y" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint8_t)' x + y" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(uint8_t)' y" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint8_t)' x + y" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 3) (TWord 3) (TWord 3)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))).
+Notation "'(uint16_t)' x + '(uint16_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 50).
+Notation "'(uint16_t)' x + '(uint16_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50).
+Notation "'(uint16_t)' x + '(uint16_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint16_t)' x + '(uint16_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(uint16_t)' y" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 50).
+Notation "'(uint16_t)' x + y" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 50).
+Notation "x + '(uint16_t)' y" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50).
+Notation "'(uint16_t)' x + y" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 50).
+Notation "x + '(uint16_t)' y" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint16_t)' x + y" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(uint16_t)' y" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint16_t)' x + y" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 4) (TWord 4) (TWord 4)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))).
+Notation "'(uint32_t)' x + '(uint32_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 50).
+Notation "'(uint32_t)' x + '(uint32_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50).
+Notation "'(uint32_t)' x + '(uint32_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint32_t)' x + '(uint32_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(uint32_t)' y" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 50).
+Notation "'(uint32_t)' x + y" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 50).
+Notation "x + '(uint32_t)' y" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50).
+Notation "'(uint32_t)' x + y" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 50).
+Notation "x + '(uint32_t)' y" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint32_t)' x + y" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(uint32_t)' y" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint32_t)' x + y" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 5) (TWord 5) (TWord 5)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))).
+Notation "'(uint64_t)' x + '(uint64_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 50).
+Notation "'(uint64_t)' x + '(uint64_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50).
+Notation "'(uint64_t)' x + '(uint64_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint64_t)' x + '(uint64_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(uint64_t)' y" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 50).
+Notation "'(uint64_t)' x + y" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 50).
+Notation "x + '(uint64_t)' y" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50).
+Notation "'(uint64_t)' x + y" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 50).
+Notation "x + '(uint64_t)' y" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint64_t)' x + y" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(uint64_t)' y" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint64_t)' x + y" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 6) (TWord 6) (TWord 6)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))).
+Notation "'(uint128_t)' x + '(uint128_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 50).
+Notation "'(uint128_t)' x + '(uint128_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50).
+Notation "'(uint128_t)' x + '(uint128_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint128_t)' x + '(uint128_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(uint128_t)' y" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 50).
+Notation "'(uint128_t)' x + y" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 50).
+Notation "x + '(uint128_t)' y" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50).
+Notation "'(uint128_t)' x + y" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 50).
+Notation "x + '(uint128_t)' y" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint128_t)' x + y" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(uint128_t)' y" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint128_t)' x + y" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 7) (TWord 7) (TWord 7)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))).
+Notation "x - y" := (Op (Sub _ _ _) (Pair x y)).
+Notation "x - y" := (Op (Sub _ _ _) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub _ _ _) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub _ _ _) (Pair (Var x) (Var y))).
+Notation "'(bool)' x - '(bool)' y" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 50).
+Notation "'(bool)' x - '(bool)' y" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50).
+Notation "'(bool)' x - '(bool)' y" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50).
+Notation "'(bool)' x - '(bool)' y" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(bool)' y" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 50).
+Notation "'(bool)' x - y" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 50).
+Notation "x - '(bool)' y" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50).
+Notation "'(bool)' x - y" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 50).
+Notation "x - '(bool)' y" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50).
+Notation "'(bool)' x - y" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(bool)' y" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(bool)' x - y" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 0) (TWord 0) (TWord 0)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x - '(uint8_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 50).
+Notation "'(uint8_t)' x - '(uint8_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50).
+Notation "'(uint8_t)' x - '(uint8_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint8_t)' x - '(uint8_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(uint8_t)' y" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 50).
+Notation "'(uint8_t)' x - y" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 50).
+Notation "x - '(uint8_t)' y" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50).
+Notation "'(uint8_t)' x - y" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 50).
+Notation "x - '(uint8_t)' y" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint8_t)' x - y" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(uint8_t)' y" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint8_t)' x - y" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 1) (TWord 1) (TWord 1)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x - '(uint8_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 50).
+Notation "'(uint8_t)' x - '(uint8_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50).
+Notation "'(uint8_t)' x - '(uint8_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint8_t)' x - '(uint8_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(uint8_t)' y" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 50).
+Notation "'(uint8_t)' x - y" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 50).
+Notation "x - '(uint8_t)' y" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50).
+Notation "'(uint8_t)' x - y" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 50).
+Notation "x - '(uint8_t)' y" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint8_t)' x - y" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(uint8_t)' y" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint8_t)' x - y" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 2) (TWord 2) (TWord 2)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x - '(uint8_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 50).
+Notation "'(uint8_t)' x - '(uint8_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50).
+Notation "'(uint8_t)' x - '(uint8_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint8_t)' x - '(uint8_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(uint8_t)' y" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 50).
+Notation "'(uint8_t)' x - y" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 50).
+Notation "x - '(uint8_t)' y" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50).
+Notation "'(uint8_t)' x - y" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 50).
+Notation "x - '(uint8_t)' y" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint8_t)' x - y" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(uint8_t)' y" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint8_t)' x - y" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 3) (TWord 3) (TWord 3)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))).
+Notation "'(uint16_t)' x - '(uint16_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 50).
+Notation "'(uint16_t)' x - '(uint16_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50).
+Notation "'(uint16_t)' x - '(uint16_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint16_t)' x - '(uint16_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(uint16_t)' y" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 50).
+Notation "'(uint16_t)' x - y" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 50).
+Notation "x - '(uint16_t)' y" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50).
+Notation "'(uint16_t)' x - y" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 50).
+Notation "x - '(uint16_t)' y" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint16_t)' x - y" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(uint16_t)' y" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint16_t)' x - y" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 4) (TWord 4) (TWord 4)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))).
+Notation "'(uint32_t)' x - '(uint32_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 50).
+Notation "'(uint32_t)' x - '(uint32_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50).
+Notation "'(uint32_t)' x - '(uint32_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint32_t)' x - '(uint32_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(uint32_t)' y" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 50).
+Notation "'(uint32_t)' x - y" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 50).
+Notation "x - '(uint32_t)' y" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50).
+Notation "'(uint32_t)' x - y" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 50).
+Notation "x - '(uint32_t)' y" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint32_t)' x - y" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(uint32_t)' y" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint32_t)' x - y" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 5) (TWord 5) (TWord 5)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))).
+Notation "'(uint64_t)' x - '(uint64_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 50).
+Notation "'(uint64_t)' x - '(uint64_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50).
+Notation "'(uint64_t)' x - '(uint64_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint64_t)' x - '(uint64_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(uint64_t)' y" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 50).
+Notation "'(uint64_t)' x - y" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 50).
+Notation "x - '(uint64_t)' y" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50).
+Notation "'(uint64_t)' x - y" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 50).
+Notation "x - '(uint64_t)' y" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint64_t)' x - y" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(uint64_t)' y" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint64_t)' x - y" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 6) (TWord 6) (TWord 6)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))).
+Notation "'(uint128_t)' x - '(uint128_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 50).
+Notation "'(uint128_t)' x - '(uint128_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50).
+Notation "'(uint128_t)' x - '(uint128_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint128_t)' x - '(uint128_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(uint128_t)' y" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 50).
+Notation "'(uint128_t)' x - y" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 50).
+Notation "x - '(uint128_t)' y" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50).
+Notation "'(uint128_t)' x - y" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 50).
+Notation "x - '(uint128_t)' y" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint128_t)' x - y" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(uint128_t)' y" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint128_t)' x - y" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 7) (TWord 7) (TWord 7)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))).
+Notation "x & y" := (Op (Land _ _ _) (Pair x y)).
+Notation "x & y" := (Op (Land _ _ _) (Pair x (Var y))).
+Notation "x & y" := (Op (Land _ _ _) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land _ _ _) (Pair (Var x) (Var y))).
+Notation "'(bool)' x & '(bool)' y" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 40).
+Notation "'(bool)' x & '(bool)' y" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40).
+Notation "'(bool)' x & '(bool)' y" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40).
+Notation "'(bool)' x & '(bool)' y" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(bool)' y" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 40).
+Notation "'(bool)' x & y" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 40).
+Notation "x & '(bool)' y" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40).
+Notation "'(bool)' x & y" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 40).
+Notation "x & '(bool)' y" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40).
+Notation "'(bool)' x & y" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(bool)' y" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(bool)' x & y" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 0) (TWord 0) (TWord 0)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x & '(uint8_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 40).
+Notation "'(uint8_t)' x & '(uint8_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40).
+Notation "'(uint8_t)' x & '(uint8_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint8_t)' x & '(uint8_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(uint8_t)' y" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 40).
+Notation "'(uint8_t)' x & y" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 40).
+Notation "x & '(uint8_t)' y" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40).
+Notation "'(uint8_t)' x & y" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 40).
+Notation "x & '(uint8_t)' y" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint8_t)' x & y" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(uint8_t)' y" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint8_t)' x & y" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 1) (TWord 1) (TWord 1)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x & '(uint8_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 40).
+Notation "'(uint8_t)' x & '(uint8_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40).
+Notation "'(uint8_t)' x & '(uint8_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint8_t)' x & '(uint8_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(uint8_t)' y" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 40).
+Notation "'(uint8_t)' x & y" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 40).
+Notation "x & '(uint8_t)' y" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40).
+Notation "'(uint8_t)' x & y" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 40).
+Notation "x & '(uint8_t)' y" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint8_t)' x & y" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(uint8_t)' y" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint8_t)' x & y" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 2) (TWord 2) (TWord 2)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x & '(uint8_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 40).
+Notation "'(uint8_t)' x & '(uint8_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40).
+Notation "'(uint8_t)' x & '(uint8_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint8_t)' x & '(uint8_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(uint8_t)' y" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 40).
+Notation "'(uint8_t)' x & y" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 40).
+Notation "x & '(uint8_t)' y" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40).
+Notation "'(uint8_t)' x & y" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 40).
+Notation "x & '(uint8_t)' y" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint8_t)' x & y" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(uint8_t)' y" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint8_t)' x & y" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 3) (TWord 3) (TWord 3)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))).
+Notation "'(uint16_t)' x & '(uint16_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 40).
+Notation "'(uint16_t)' x & '(uint16_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40).
+Notation "'(uint16_t)' x & '(uint16_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint16_t)' x & '(uint16_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(uint16_t)' y" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 40).
+Notation "'(uint16_t)' x & y" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 40).
+Notation "x & '(uint16_t)' y" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40).
+Notation "'(uint16_t)' x & y" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 40).
+Notation "x & '(uint16_t)' y" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint16_t)' x & y" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(uint16_t)' y" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint16_t)' x & y" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 4) (TWord 4) (TWord 4)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))).
+Notation "'(uint32_t)' x & '(uint32_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 40).
+Notation "'(uint32_t)' x & '(uint32_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40).
+Notation "'(uint32_t)' x & '(uint32_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint32_t)' x & '(uint32_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(uint32_t)' y" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 40).
+Notation "'(uint32_t)' x & y" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 40).
+Notation "x & '(uint32_t)' y" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40).
+Notation "'(uint32_t)' x & y" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 40).
+Notation "x & '(uint32_t)' y" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint32_t)' x & y" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(uint32_t)' y" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint32_t)' x & y" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 5) (TWord 5) (TWord 5)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))).
+Notation "'(uint64_t)' x & '(uint64_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 40).
+Notation "'(uint64_t)' x & '(uint64_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40).
+Notation "'(uint64_t)' x & '(uint64_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint64_t)' x & '(uint64_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(uint64_t)' y" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 40).
+Notation "'(uint64_t)' x & y" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 40).
+Notation "x & '(uint64_t)' y" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40).
+Notation "'(uint64_t)' x & y" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 40).
+Notation "x & '(uint64_t)' y" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint64_t)' x & y" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(uint64_t)' y" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint64_t)' x & y" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 6) (TWord 6) (TWord 6)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))).
+Notation "'(uint128_t)' x & '(uint128_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 40).
+Notation "'(uint128_t)' x & '(uint128_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40).
+Notation "'(uint128_t)' x & '(uint128_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint128_t)' x & '(uint128_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(uint128_t)' y" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 40).
+Notation "'(uint128_t)' x & y" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 40).
+Notation "x & '(uint128_t)' y" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40).
+Notation "'(uint128_t)' x & y" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 40).
+Notation "x & '(uint128_t)' y" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint128_t)' x & y" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(uint128_t)' y" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint128_t)' x & y" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 7) (TWord 7) (TWord 7)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))).
+Notation "x << y" := (Op (Shl _ _ _) (Pair x y)).
+Notation "x << y" := (Op (Shl _ _ _) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl _ _ _) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl _ _ _) (Pair (Var x) (Var y))).
+Notation "'(bool)' x << '(bool)' y" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 30).
+Notation "'(bool)' x << '(bool)' y" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 30).
+Notation "'(bool)' x << '(bool)' y" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 30).
+Notation "'(bool)' x << '(bool)' y" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(bool)' y" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 30).
+Notation "'(bool)' x << y" := (Op (Shl (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 30).
+Notation "x << '(bool)' y" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 30).
+Notation "'(bool)' x << y" := (Op (Shl (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 30).
+Notation "x << '(bool)' y" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 30).
+Notation "'(bool)' x << y" := (Op (Shl (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(bool)' y" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(bool)' x << y" := (Op (Shl (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 0) (TWord 0) (TWord 0)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x << '(uint8_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 30).
+Notation "'(uint8_t)' x << '(uint8_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 30).
+Notation "'(uint8_t)' x << '(uint8_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint8_t)' x << '(uint8_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(uint8_t)' y" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 30).
+Notation "'(uint8_t)' x << y" := (Op (Shl (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 30).
+Notation "x << '(uint8_t)' y" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 30).
+Notation "'(uint8_t)' x << y" := (Op (Shl (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 30).
+Notation "x << '(uint8_t)' y" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint8_t)' x << y" := (Op (Shl (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(uint8_t)' y" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint8_t)' x << y" := (Op (Shl (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 1) (TWord 1) (TWord 1)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x << '(uint8_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 30).
+Notation "'(uint8_t)' x << '(uint8_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 30).
+Notation "'(uint8_t)' x << '(uint8_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint8_t)' x << '(uint8_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(uint8_t)' y" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 30).
+Notation "'(uint8_t)' x << y" := (Op (Shl (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 30).
+Notation "x << '(uint8_t)' y" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 30).
+Notation "'(uint8_t)' x << y" := (Op (Shl (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 30).
+Notation "x << '(uint8_t)' y" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint8_t)' x << y" := (Op (Shl (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(uint8_t)' y" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint8_t)' x << y" := (Op (Shl (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 2) (TWord 2) (TWord 2)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))).
+Notation "'(uint8_t)' x << '(uint8_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 30).
+Notation "'(uint8_t)' x << '(uint8_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 30).
+Notation "'(uint8_t)' x << '(uint8_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint8_t)' x << '(uint8_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(uint8_t)' y" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 30).
+Notation "'(uint8_t)' x << y" := (Op (Shl (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 30).
+Notation "x << '(uint8_t)' y" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 30).
+Notation "'(uint8_t)' x << y" := (Op (Shl (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 30).
+Notation "x << '(uint8_t)' y" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint8_t)' x << y" := (Op (Shl (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(uint8_t)' y" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint8_t)' x << y" := (Op (Shl (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 3) (TWord 3) (TWord 3)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))).
+Notation "'(uint16_t)' x << '(uint16_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 30).
+Notation "'(uint16_t)' x << '(uint16_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 30).
+Notation "'(uint16_t)' x << '(uint16_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint16_t)' x << '(uint16_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(uint16_t)' y" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 30).
+Notation "'(uint16_t)' x << y" := (Op (Shl (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 30).
+Notation "x << '(uint16_t)' y" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 30).
+Notation "'(uint16_t)' x << y" := (Op (Shl (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 30).
+Notation "x << '(uint16_t)' y" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint16_t)' x << y" := (Op (Shl (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(uint16_t)' y" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint16_t)' x << y" := (Op (Shl (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 4) (TWord 4) (TWord 4)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))).
+Notation "'(uint32_t)' x << '(uint32_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 30).
+Notation "'(uint32_t)' x << '(uint32_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 30).
+Notation "'(uint32_t)' x << '(uint32_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint32_t)' x << '(uint32_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(uint32_t)' y" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 30).
+Notation "'(uint32_t)' x << y" := (Op (Shl (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 30).
+Notation "x << '(uint32_t)' y" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 30).
+Notation "'(uint32_t)' x << y" := (Op (Shl (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 30).
+Notation "x << '(uint32_t)' y" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint32_t)' x << y" := (Op (Shl (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(uint32_t)' y" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint32_t)' x << y" := (Op (Shl (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 5) (TWord 5) (TWord 5)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))).
+Notation "'(uint64_t)' x << '(uint64_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 30).
+Notation "'(uint64_t)' x << '(uint64_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 30).
+Notation "'(uint64_t)' x << '(uint64_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint64_t)' x << '(uint64_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(uint64_t)' y" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 30).
+Notation "'(uint64_t)' x << y" := (Op (Shl (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 30).
+Notation "x << '(uint64_t)' y" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 30).
+Notation "'(uint64_t)' x << y" := (Op (Shl (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 30).
+Notation "x << '(uint64_t)' y" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint64_t)' x << y" := (Op (Shl (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(uint64_t)' y" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint64_t)' x << y" := (Op (Shl (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 6) (TWord 6) (TWord 6)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))).
+Notation "'(uint128_t)' x << '(uint128_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 30).
+Notation "'(uint128_t)' x << '(uint128_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 30).
+Notation "'(uint128_t)' x << '(uint128_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint128_t)' x << '(uint128_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(uint128_t)' y" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 30).
+Notation "'(uint128_t)' x << y" := (Op (Shl (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 30).
+Notation "x << '(uint128_t)' y" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 30).
+Notation "'(uint128_t)' x << y" := (Op (Shl (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 30).
+Notation "x << '(uint128_t)' y" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint128_t)' x << y" := (Op (Shl (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(uint128_t)' y" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint128_t)' x << y" := (Op (Shl (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 7) (TWord 7) (TWord 7)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))).
+Notation "x >> y" := (Op (Shr _ _ _) (Pair x y)).
+Notation "x >> y" := (Op (Shr _ _ _) (Pair x (Var y))).
+Notation "x >> y" := (Op (Shr _ _ _) (Pair (Var x) y)).
+Notation "x >> y" := (Op (Shr _ _ _) (Pair (Var x) (Var y))).
+Notation "'(bool)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 30).
+Notation "'(bool)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 30).
+Notation "'(bool)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 30).
+Notation "'(bool)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 30).
+Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 30).
+Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 30).
+Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 30).
+Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 30).
+Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 30).
+Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint16_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 30).
+Notation "'(uint16_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 30).
+Notation "'(uint16_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint16_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint32_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 30).
+Notation "'(uint32_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 30).
+Notation "'(uint32_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint32_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint64_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 30).
+Notation "'(uint64_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 30).
+Notation "'(uint64_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint64_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint128_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 30).
+Notation "'(uint128_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 30).
+Notation "'(uint128_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint128_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 30).
+Notation Return x := (Var x).
+Notation C_like := (Expr base_type op _).
diff --git a/src/Compilers/Z/FoldTypes.v b/src/Compilers/Z/FoldTypes.v
new file mode 100644
index 000000000..fc6aa9406
--- /dev/null
+++ b/src/Compilers/Z/FoldTypes.v
@@ -0,0 +1,17 @@
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.Util.
+Require Import Crypto.Compilers.FoldTypes.
+
+Section min_or_max.
+ Context (f : base_type -> base_type -> base_type)
+ (init : base_type).
+
+ Definition TypeFold {t} (e : Expr base_type op t) : base_type
+ := TypeFold (fun t => t) f init e.
+End min_or_max.
+
+Definition MaxTypeUsed {t} (e : Expr base_type op t) : base_type
+ := TypeFold base_type_max (TWord 0) e.
+Definition MinTypeUsed {t} (e : Expr base_type op t) : base_type
+ := TypeFold base_type_min TZ e.
diff --git a/src/Compilers/Z/HexNotationConstants.v b/src/Compilers/Z/HexNotationConstants.v
new file mode 100644
index 000000000..a5fb9b757
--- /dev/null
+++ b/src/Compilers/Z/HexNotationConstants.v
@@ -0,0 +1,144 @@
+Require Import Coq.ZArith.ZArith.
+Require Export Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Export Bedrock.Word.
+Require Export Crypto.Util.Notations.
+
+Notation Const x := (Op (OpConst x) TT).
+(* python:
+<<
+print('\n'.join('''Notation "'%s'" (* %d (%s) *)\n := (Const %s%%Z).\nNotation "'%s'" (* %d (%s) *)\n := (Const %s).''' % (h, d, h, d, h, d, h, w)
+ for d, h, b, w in sorted([(eval(bv), hex(eval(bv)), bv, i)
+ for (bv, i) in (('0b' + i[2:].replace('~', ''), i)
+ for i in r"""WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1
+WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~1~0
+WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0
+WO~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1
+WO~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0
+WO~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1
+WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~1~0
+WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~0
+WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0
+WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1
+WO~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0
+WO~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1
+WO~0~0~0~1~0~0~1~1
+WO~0~0~0~1~1~0~0~1
+WO~0~0~0~1~1~0~1~0
+WO~0~0~0~1~1~0~1~1
+WO~0~0~0~1~1~1~0~0
+WO~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0
+WO~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0
+WO~0~0~1~1~0~0~1~1
+WO~1~0
+WO~1~0~0~1
+WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1
+WO~0~0~0~1~0~0~0~1
+WO~0~0~0~1~0~1~1~1
+WO~1~1""".split('\n'))])))
+>>
+ *)
+Notation "'0x2'" (* 2 (0x2) *)
+ := (Const 2%Z).
+Notation "'0x2'" (* 2 (0x2) *)
+ := (Const WO~1~0).
+Notation "'0x3'" (* 3 (0x3) *)
+ := (Const 3%Z).
+Notation "'0x3'" (* 3 (0x3) *)
+ := (Const WO~1~1).
+Notation "'0x9'" (* 9 (0x9) *)
+ := (Const 9%Z).
+Notation "'0x9'" (* 9 (0x9) *)
+ := (Const WO~1~0~0~1).
+Notation "'0x11'" (* 17 (0x11) *)
+ := (Const 17%Z).
+Notation "'0x11'" (* 17 (0x11) *)
+ := (Const WO~0~0~0~1~0~0~0~1).
+Notation "'0x13'" (* 19 (0x13) *)
+ := (Const 19%Z).
+Notation "'0x13'" (* 19 (0x13) *)
+ := (Const WO~0~0~0~1~0~0~1~1).
+Notation "'0x17'" (* 23 (0x17) *)
+ := (Const 23%Z).
+Notation "'0x17'" (* 23 (0x17) *)
+ := (Const WO~0~0~0~1~0~1~1~1).
+Notation "'0x19'" (* 25 (0x19) *)
+ := (Const 25%Z).
+Notation "'0x19'" (* 25 (0x19) *)
+ := (Const WO~0~0~0~1~1~0~0~1).
+Notation "'0x1a'" (* 26 (0x1a) *)
+ := (Const 26%Z).
+Notation "'0x1a'" (* 26 (0x1a) *)
+ := (Const WO~0~0~0~1~1~0~1~0).
+Notation "'0x1b'" (* 27 (0x1b) *)
+ := (Const 27%Z).
+Notation "'0x1b'" (* 27 (0x1b) *)
+ := (Const WO~0~0~0~1~1~0~1~1).
+Notation "'0x1c'" (* 28 (0x1c) *)
+ := (Const 28%Z).
+Notation "'0x1c'" (* 28 (0x1c) *)
+ := (Const WO~0~0~0~1~1~1~0~0).
+Notation "'0x33'" (* 51 (0x33) *)
+ := (Const 51%Z).
+Notation "'0x33'" (* 51 (0x33) *)
+ := (Const WO~0~0~1~1~0~0~1~1).
+Notation "'0x7fffff'" (* 8388607 (0x7fffff) *)
+ := (Const 8388607%Z).
+Notation "'0x7fffff'" (* 8388607 (0x7fffff) *)
+ := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
+Notation "'0x1ffffff'" (* 33554431 (0x1ffffff) *)
+ := (Const 33554431%Z).
+Notation "'0x1ffffff'" (* 33554431 (0x1ffffff) *)
+ := (Const WO~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
+Notation "'0x3fffffe'" (* 67108862 (0x3fffffe) *)
+ := (Const 67108862%Z).
+Notation "'0x3fffffe'" (* 67108862 (0x3fffffe) *)
+ := (Const WO~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
+Notation "'0x3ffffff'" (* 67108863 (0x3ffffff) *)
+ := (Const 67108863%Z).
+Notation "'0x3ffffff'" (* 67108863 (0x3ffffff) *)
+ := (Const WO~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
+Notation "'0x7ffffda'" (* 134217690 (0x7ffffda) *)
+ := (Const 134217690%Z).
+Notation "'0x7ffffda'" (* 134217690 (0x7ffffda) *)
+ := (Const WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~1~0).
+Notation "'0x7ffffee'" (* 134217710 (0x7ffffee) *)
+ := (Const 134217710%Z).
+Notation "'0x7ffffee'" (* 134217710 (0x7ffffee) *)
+ := (Const WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~0).
+Notation "'0x7fffffe'" (* 134217726 (0x7fffffe) *)
+ := (Const 134217726%Z).
+Notation "'0x7fffffe'" (* 134217726 (0x7fffffe) *)
+ := (Const WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
+Notation "'0x7ffffff'" (* 134217727 (0x7ffffff) *)
+ := (Const 134217727%Z).
+Notation "'0x7ffffff'" (* 134217727 (0x7ffffff) *)
+ := (Const WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
+Notation "'0xffffffe'" (* 268435454 (0xffffffe) *)
+ := (Const 268435454%Z).
+Notation "'0xffffffe'" (* 268435454 (0xffffffe) *)
+ := (Const WO~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
+Notation "'0xfffffff'" (* 268435455 (0xfffffff) *)
+ := (Const 268435455%Z).
+Notation "'0xfffffff'" (* 268435455 (0xfffffff) *)
+ := (Const WO~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
+Notation "'0x1ffffffa'" (* 536870906 (0x1ffffffa) *)
+ := (Const 536870906%Z).
+Notation "'0x1ffffffa'" (* 536870906 (0x1ffffffa) *)
+ := (Const WO~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0).
+Notation "'0x1ffffffe'" (* 536870910 (0x1ffffffe) *)
+ := (Const 536870910%Z).
+Notation "'0x1ffffffe'" (* 536870910 (0x1ffffffe) *)
+ := (Const WO~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
+Notation "'0x7ffffffffffff'" (* 2251799813685247 (0x7ffffffffffff) *)
+ := (Const 2251799813685247%Z).
+Notation "'0x7ffffffffffff'" (* 2251799813685247 (0x7ffffffffffff) *)
+ := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
+Notation "'0xfffffffffffda'" (* 4503599627370458 (0xfffffffffffda) *)
+ := (Const 4503599627370458%Z).
+Notation "'0xfffffffffffda'" (* 4503599627370458 (0xfffffffffffda) *)
+ := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~1~0).
+Notation "'0xffffffffffffe'" (* 4503599627370494 (0xffffffffffffe) *)
+ := (Const 4503599627370494%Z).
+Notation "'0xffffffffffffe'" (* 4503599627370494 (0xffffffffffffe) *)
+ := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
diff --git a/src/Compilers/Z/Inline.v b/src/Compilers/Z/Inline.v
new file mode 100644
index 000000000..8a8bea98b
--- /dev/null
+++ b/src/Compilers/Z/Inline.v
@@ -0,0 +1,7 @@
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Inline.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.Util.
+
+Definition InlineConst {t} (e : Expr base_type op t) : Expr base_type op t
+ := @InlineConst base_type op (is_const) t e.
diff --git a/src/Compilers/Z/InlineInterp.v b/src/Compilers/Z/InlineInterp.v
new file mode 100644
index 000000000..c02197331
--- /dev/null
+++ b/src/Compilers/Z/InlineInterp.v
@@ -0,0 +1,11 @@
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Wf.
+Require Import Crypto.Compilers.InlineInterp.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Inline.
+
+Definition InterpInlineConst {interp_base_type interp_op} {t} (e : Expr base_type op t) (Hwf : Wf e)
+ : forall x, Interp interp_op (InlineConst e) x = Interp interp_op e x
+ := @InterpInlineConst _ interp_base_type _ _ _ t e Hwf.
+
+Hint Rewrite @InterpInlineConst using solve [ eassumption | eauto with wf ] : reflective_interp.
diff --git a/src/Compilers/Z/InlineWf.v b/src/Compilers/Z/InlineWf.v
new file mode 100644
index 000000000..32b84aa1f
--- /dev/null
+++ b/src/Compilers/Z/InlineWf.v
@@ -0,0 +1,11 @@
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Wf.
+Require Import Crypto.Compilers.InlineWf.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Inline.
+
+Definition Wf_InlineConst {t} (e : Expr base_type op t) (Hwf : Wf e)
+ : Wf (InlineConst e)
+ := @Wf_InlineConst _ _ _ t e Hwf.
+
+Hint Resolve Wf_InlineConst : wf.
diff --git a/src/Compilers/Z/JavaNotations.v b/src/Compilers/Z/JavaNotations.v
new file mode 100644
index 000000000..bab120b0a
--- /dev/null
+++ b/src/Compilers/Z/JavaNotations.v
@@ -0,0 +1,792 @@
+Require Export Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Export Crypto.Compilers.Z.HexNotationConstants.
+Require Export Crypto.Util.Notations.
+
+Reserved Notation "T x = A ; b" (at level 200, b at level 200, format "T x = A ; '//' b").
+Reserved Notation "x & y" (at level 40).
+(* N.B. M32 is 0xFFFFFFFFL, and is how to cast a 64-bit thing to a 32-bit thing in Java *)
+Reserved Notation "'M32' & x" (at level 200, x at level 9).
+
+Global Open Scope expr_scope.
+
+Notation "T x = A ; b" := (LetIn (tx:=T) A (fun x => b)) : expr_scope.
+(* ??? Did I get M32 wrong? *)
+(*Notation "'(int)' x" := (Op (Cast _ (TWord 0)) x).
+Notation "'(int)' x" := (Op (Cast _ (TWord 1)) x).
+Notation "'(int)' x" := (Op (Cast _ (TWord 2)) x).
+Notation "'(int)' x" := (Op (Cast _ (TWord 3)) x).
+Notation "'(int)' x" := (Op (Cast _ (TWord 4)) x).
+Notation "'(int)' x" := (Op (Cast _ (TWord 5)) x).
+Notation "'M32' & x" := (Op (Cast _ (TWord 6)) x).
+Notation "'(uint128_t)' x" := (Op (Cast _ (TWord 7)) x).
+Notation "'(int)' x" := (Op (Cast _ (TWord 0)) (Var x)).
+Notation "'(int)' x" := (Op (Cast _ (TWord 1)) (Var x)).
+Notation "'(int)' x" := (Op (Cast _ (TWord 2)) (Var x)).
+Notation "'(int)' x" := (Op (Cast _ (TWord 3)) (Var x)).
+Notation "'(int)' x" := (Op (Cast _ (TWord 4)) (Var x)).
+Notation "'(int)' x" := (Op (Cast _ (TWord 5)) (Var x)).
+Notation "'M32' & x" := (Op (Cast _ (TWord 6)) (Var x)).
+Notation "'(uint128_t)' x" := (Op (Cast _ (TWord 7)) (Var x)).*)
+(* python:
+<<
+types = ('int', 'int', 'int', 'int', 'int', 'int', 'long', 'uint128_t')
+for lgwordsz in range(0, len(types)):
+ print('Notation "\'%s\'" := (Tbase (TWord %d)).' % (types[lgwordsz], lgwordsz))
+print('Notation ℤ := (Tbase TZ).')
+print('')
+cast_pat = "'(%s)' %s"
+for opn, op, lvl in (('*', 'Mul', 40), ('+', 'Add', 50), ('-', 'Sub', 50), ('&', 'Land', 40), ('<<', 'Shl', 30)):
+ for v1 in (False, True):
+ for v2 in (False, True):
+ lhs = ('x' if not v1 else '(Var x)')
+ rhs = ('y' if not v2 else '(Var y)')
+ print('Notation "x %s y" := (Op (%s _ _ _) (Pair %s %s)).' % (opn, op, lhs, rhs))
+ for lgwordsz in range(0, len(types)):
+ for v1 in (False, True):
+ for v2 in (False, True):
+ lhs = ('x' if not v1 else '(Var x)')
+ rhs = ('y' if not v2 else '(Var y)')
+ print('Notation "%s %s %s" := (Op (%s (TWord _) (TWord _) (TWord %d)) (Pair %s %s)) (at level %d).'
+ % (cast_pat % (types[lgwordsz], 'x'), opn, cast_pat % (types[lgwordsz], 'y'),
+ op, lgwordsz, lhs, rhs, lvl))
+ for v1 in (False, True):
+ for v2 in (False, True):
+ lhs = ('x' if not v1 else '(Var x)')
+ rhs = ('y' if not v2 else '(Var y)')
+ print('Notation "%s %s %s" := (Op (%s (TWord %d) (TWord _) (TWord %d)) (Pair %s %s)) (at level %d).'
+ % ('x', opn, cast_pat % (types[lgwordsz], 'y'),
+ op, lgwordsz, lgwordsz, lhs, rhs, lvl))
+ print('Notation "%s %s %s" := (Op (%s (TWord _) (TWord %d) (TWord %d)) (Pair %s %s)) (at level %d).'
+ % (cast_pat % (types[lgwordsz], 'x'), opn, 'y',
+ op, lgwordsz, lgwordsz, lhs, rhs, lvl))
+ for v1 in (False, True):
+ for v2 in (False, True):
+ lhs = ('x' if not v1 else '(Var x)')
+ rhs = ('y' if not v2 else '(Var y)')
+ print('Notation "x %s y" := (Op (%s (TWord %d) (TWord %d) (TWord %d)) (Pair %s %s)).'
+ % (opn, op, lgwordsz, lgwordsz, lgwordsz, lhs, rhs))
+for opn, op, lvl in (('>>', 'Shr', 30),):
+ for v1 in (False, True):
+ for v2 in (False, True):
+ lhs = ('x' if not v1 else '(Var x)')
+ rhs = ('y' if not v2 else '(Var y)')
+ print('Notation "x %s y" := (Op (%s _ _ _) (Pair %s %s)).' % (opn, op, lhs, rhs))
+ for lgwordsz in range(0, len(types)):
+ for v1 in (False, True):
+ for v2 in (False, True):
+ lhs = ('x' if not v1 else '(Var x)')
+ rhs = ('y' if not v2 else '(Var y)')
+ print('Notation "\'(%s)\' ( x %s y )" := (Op (%s (TWord _) (TWord _) (TWord %d)) (Pair %s %s)) (at level %d).'
+ % (types[lgwordsz], opn, op, lgwordsz, lhs, rhs, lvl))
+print('Notation Return x := (Var x).')
+print('Notation Java_like := (Expr base_type op _).')
+>> *)
+Notation "'int'" := (Tbase (TWord 0)).
+Notation "'int'" := (Tbase (TWord 1)).
+Notation "'int'" := (Tbase (TWord 2)).
+Notation "'int'" := (Tbase (TWord 3)).
+Notation "'int'" := (Tbase (TWord 4)).
+Notation "'int'" := (Tbase (TWord 5)).
+Notation "'long'" := (Tbase (TWord 6)).
+Notation "'uint128_t'" := (Tbase (TWord 7)).
+Notation ℤ := (Tbase TZ).
+
+Notation "x * y" := (Op (Mul _ _ _) (Pair x y)).
+Notation "x * y" := (Op (Mul _ _ _) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul _ _ _) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul _ _ _) (Pair (Var x) (Var y))).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 0) (TWord 0) (TWord 0)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 1) (TWord 1) (TWord 1)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 2) (TWord 2) (TWord 2)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 3) (TWord 3) (TWord 3)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 4) (TWord 4) (TWord 4)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(int)' y" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 5) (TWord 5) (TWord 5)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))).
+Notation "'(long)' x * '(long)' y" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 40).
+Notation "'(long)' x * '(long)' y" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40).
+Notation "'(long)' x * '(long)' y" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40).
+Notation "'(long)' x * '(long)' y" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(long)' y" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 40).
+Notation "'(long)' x * y" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 40).
+Notation "x * '(long)' y" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40).
+Notation "'(long)' x * y" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 40).
+Notation "x * '(long)' y" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40).
+Notation "'(long)' x * y" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(long)' y" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(long)' x * y" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 6) (TWord 6) (TWord 6)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))).
+Notation "'(uint128_t)' x * '(uint128_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 40).
+Notation "'(uint128_t)' x * '(uint128_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40).
+Notation "'(uint128_t)' x * '(uint128_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint128_t)' x * '(uint128_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * '(uint128_t)' y" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 40).
+Notation "'(uint128_t)' x * y" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 40).
+Notation "x * '(uint128_t)' y" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40).
+Notation "'(uint128_t)' x * y" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 40).
+Notation "x * '(uint128_t)' y" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint128_t)' x * y" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 40).
+Notation "x * '(uint128_t)' y" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint128_t)' x * y" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x * y" := (Op (Mul (TWord 7) (TWord 7) (TWord 7)) (Pair x y)).
+Notation "x * y" := (Op (Mul (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))).
+Notation "x * y" := (Op (Mul (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)).
+Notation "x * y" := (Op (Mul (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))).
+Notation "x + y" := (Op (Add _ _ _) (Pair x y)).
+Notation "x + y" := (Op (Add _ _ _) (Pair x (Var y))).
+Notation "x + y" := (Op (Add _ _ _) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add _ _ _) (Pair (Var x) (Var y))).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 0) (TWord 0) (TWord 0)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 1) (TWord 1) (TWord 1)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 2) (TWord 2) (TWord 2)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 3) (TWord 3) (TWord 3)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 4) (TWord 4) (TWord 4)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(int)' y" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 5) (TWord 5) (TWord 5)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))).
+Notation "'(long)' x + '(long)' y" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 50).
+Notation "'(long)' x + '(long)' y" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50).
+Notation "'(long)' x + '(long)' y" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50).
+Notation "'(long)' x + '(long)' y" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(long)' y" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 50).
+Notation "'(long)' x + y" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 50).
+Notation "x + '(long)' y" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50).
+Notation "'(long)' x + y" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 50).
+Notation "x + '(long)' y" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50).
+Notation "'(long)' x + y" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(long)' y" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(long)' x + y" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 6) (TWord 6) (TWord 6)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))).
+Notation "'(uint128_t)' x + '(uint128_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 50).
+Notation "'(uint128_t)' x + '(uint128_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50).
+Notation "'(uint128_t)' x + '(uint128_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint128_t)' x + '(uint128_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + '(uint128_t)' y" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 50).
+Notation "'(uint128_t)' x + y" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 50).
+Notation "x + '(uint128_t)' y" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50).
+Notation "'(uint128_t)' x + y" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 50).
+Notation "x + '(uint128_t)' y" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint128_t)' x + y" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 50).
+Notation "x + '(uint128_t)' y" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint128_t)' x + y" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x + y" := (Op (Add (TWord 7) (TWord 7) (TWord 7)) (Pair x y)).
+Notation "x + y" := (Op (Add (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))).
+Notation "x + y" := (Op (Add (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)).
+Notation "x + y" := (Op (Add (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))).
+Notation "x - y" := (Op (Sub _ _ _) (Pair x y)).
+Notation "x - y" := (Op (Sub _ _ _) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub _ _ _) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub _ _ _) (Pair (Var x) (Var y))).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 0) (TWord 0) (TWord 0)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 1) (TWord 1) (TWord 1)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 2) (TWord 2) (TWord 2)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 3) (TWord 3) (TWord 3)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 4) (TWord 4) (TWord 4)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(int)' y" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 5) (TWord 5) (TWord 5)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))).
+Notation "'(long)' x - '(long)' y" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 50).
+Notation "'(long)' x - '(long)' y" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50).
+Notation "'(long)' x - '(long)' y" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50).
+Notation "'(long)' x - '(long)' y" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(long)' y" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 50).
+Notation "'(long)' x - y" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 50).
+Notation "x - '(long)' y" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50).
+Notation "'(long)' x - y" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 50).
+Notation "x - '(long)' y" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50).
+Notation "'(long)' x - y" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(long)' y" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(long)' x - y" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 6) (TWord 6) (TWord 6)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))).
+Notation "'(uint128_t)' x - '(uint128_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 50).
+Notation "'(uint128_t)' x - '(uint128_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50).
+Notation "'(uint128_t)' x - '(uint128_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint128_t)' x - '(uint128_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - '(uint128_t)' y" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 50).
+Notation "'(uint128_t)' x - y" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 50).
+Notation "x - '(uint128_t)' y" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50).
+Notation "'(uint128_t)' x - y" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 50).
+Notation "x - '(uint128_t)' y" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50).
+Notation "'(uint128_t)' x - y" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 50).
+Notation "x - '(uint128_t)' y" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50).
+Notation "'(uint128_t)' x - y" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 50).
+Notation "x - y" := (Op (Sub (TWord 7) (TWord 7) (TWord 7)) (Pair x y)).
+Notation "x - y" := (Op (Sub (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))).
+Notation "x - y" := (Op (Sub (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)).
+Notation "x - y" := (Op (Sub (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))).
+Notation "x & y" := (Op (Land _ _ _) (Pair x y)).
+Notation "x & y" := (Op (Land _ _ _) (Pair x (Var y))).
+Notation "x & y" := (Op (Land _ _ _) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land _ _ _) (Pair (Var x) (Var y))).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 0) (TWord 0) (TWord 0)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 1) (TWord 1) (TWord 1)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 2) (TWord 2) (TWord 2)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 3) (TWord 3) (TWord 3)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 4) (TWord 4) (TWord 4)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(int)' y" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 5) (TWord 5) (TWord 5)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))).
+Notation "'(long)' x & '(long)' y" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 40).
+Notation "'(long)' x & '(long)' y" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40).
+Notation "'(long)' x & '(long)' y" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40).
+Notation "'(long)' x & '(long)' y" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(long)' y" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 40).
+Notation "'(long)' x & y" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 40).
+Notation "x & '(long)' y" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40).
+Notation "'(long)' x & y" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 40).
+Notation "x & '(long)' y" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40).
+Notation "'(long)' x & y" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(long)' y" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(long)' x & y" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 6) (TWord 6) (TWord 6)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))).
+Notation "'(uint128_t)' x & '(uint128_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 40).
+Notation "'(uint128_t)' x & '(uint128_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40).
+Notation "'(uint128_t)' x & '(uint128_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint128_t)' x & '(uint128_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & '(uint128_t)' y" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 40).
+Notation "'(uint128_t)' x & y" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 40).
+Notation "x & '(uint128_t)' y" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40).
+Notation "'(uint128_t)' x & y" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 40).
+Notation "x & '(uint128_t)' y" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40).
+Notation "'(uint128_t)' x & y" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 40).
+Notation "x & '(uint128_t)' y" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40).
+Notation "'(uint128_t)' x & y" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 40).
+Notation "x & y" := (Op (Land (TWord 7) (TWord 7) (TWord 7)) (Pair x y)).
+Notation "x & y" := (Op (Land (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))).
+Notation "x & y" := (Op (Land (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)).
+Notation "x & y" := (Op (Land (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))).
+Notation "x << y" := (Op (Shl _ _ _) (Pair x y)).
+Notation "x << y" := (Op (Shl _ _ _) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl _ _ _) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl _ _ _) (Pair (Var x) (Var y))).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 0) (TWord 0) (TWord 0)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 1) (TWord 1) (TWord 1)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 2) (TWord 2) (TWord 2)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 3) (TWord 3) (TWord 3)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 4) (TWord 4) (TWord 4)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(int)' y" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 5) (TWord 5) (TWord 5)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))).
+Notation "'(long)' x << '(long)' y" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 30).
+Notation "'(long)' x << '(long)' y" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 30).
+Notation "'(long)' x << '(long)' y" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 30).
+Notation "'(long)' x << '(long)' y" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(long)' y" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 30).
+Notation "'(long)' x << y" := (Op (Shl (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 30).
+Notation "x << '(long)' y" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 30).
+Notation "'(long)' x << y" := (Op (Shl (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 30).
+Notation "x << '(long)' y" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 30).
+Notation "'(long)' x << y" := (Op (Shl (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(long)' y" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(long)' x << y" := (Op (Shl (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 6) (TWord 6) (TWord 6)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))).
+Notation "'(uint128_t)' x << '(uint128_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 30).
+Notation "'(uint128_t)' x << '(uint128_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 30).
+Notation "'(uint128_t)' x << '(uint128_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint128_t)' x << '(uint128_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << '(uint128_t)' y" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 30).
+Notation "'(uint128_t)' x << y" := (Op (Shl (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 30).
+Notation "x << '(uint128_t)' y" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 30).
+Notation "'(uint128_t)' x << y" := (Op (Shl (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 30).
+Notation "x << '(uint128_t)' y" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint128_t)' x << y" := (Op (Shl (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 30).
+Notation "x << '(uint128_t)' y" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint128_t)' x << y" := (Op (Shl (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 30).
+Notation "x << y" := (Op (Shl (TWord 7) (TWord 7) (TWord 7)) (Pair x y)).
+Notation "x << y" := (Op (Shl (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))).
+Notation "x << y" := (Op (Shl (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)).
+Notation "x << y" := (Op (Shl (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))).
+Notation "x >> y" := (Op (Shr _ _ _) (Pair x y)).
+Notation "x >> y" := (Op (Shr _ _ _) (Pair x (Var y))).
+Notation "x >> y" := (Op (Shr _ _ _) (Pair (Var x) y)).
+Notation "x >> y" := (Op (Shr _ _ _) (Pair (Var x) (Var y))).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 30).
+Notation "'(int)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(long)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 30).
+Notation "'(long)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 30).
+Notation "'(long)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 30).
+Notation "'(long)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 30).
+Notation "'(uint128_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 30).
+Notation "'(uint128_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 30).
+Notation "'(uint128_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 30).
+Notation "'(uint128_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 30).
+Notation Return x := (Var x).
+Notation Java_like := (Expr base_type op _).
diff --git a/src/Compilers/Z/MapCastByDeBruijn.v b/src/Compilers/Z/MapCastByDeBruijn.v
new file mode 100644
index 000000000..1985653d4
--- /dev/null
+++ b/src/Compilers/Z/MapCastByDeBruijn.v
@@ -0,0 +1,28 @@
+Require Import Coq.ZArith.ZArith.
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.SmartMap.
+Require Import Crypto.Compilers.MapCastByDeBruijn.
+Require Import Crypto.Compilers.Z.Syntax.
+
+Section language.
+ Context {interp_base_type_bounds : base_type -> Type}
+ (interp_op_bounds : forall src dst, op src dst -> interp_flat_type interp_base_type_bounds src -> interp_flat_type interp_base_type_bounds dst)
+ (pick_typeb : forall t, interp_base_type_bounds t -> base_type).
+ Local Notation pick_type v := (SmartFlatTypeMap pick_typeb v).
+ Context (cast_op : forall t tR (opc : op t tR) args_bs,
+ op (pick_type args_bs) (pick_type (interp_op_bounds t tR opc args_bs))).
+ Context {t : type base_type}.
+
+ Definition MapCastCompile := @MapCastCompile base_type op t.
+ Definition MapCastDoCast
+ := @MapCastDoCast
+ base_type op base_type_beq internal_base_type_dec_bl
+ interp_base_type_bounds interp_op_bounds pick_typeb cast_op t.
+ Definition MapCastDoInterp
+ := @MapCastDoInterp
+ base_type op base_type_beq internal_base_type_dec_bl
+ (fun _ t => Op (OpConst 0%Z) TT)
+ interp_base_type_bounds pick_typeb t.
+ Definition MapCast e input_bounds
+ := MapCastDoInterp input_bounds (MapCastDoCast input_bounds (MapCastCompile e)).
+End language.
diff --git a/src/Compilers/Z/MapCastByDeBruijnInterp.v b/src/Compilers/Z/MapCastByDeBruijnInterp.v
new file mode 100644
index 000000000..cd145c6fa
--- /dev/null
+++ b/src/Compilers/Z/MapCastByDeBruijnInterp.v
@@ -0,0 +1,50 @@
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Wf.
+Require Import Crypto.Compilers.SmartMap.
+Require Import Crypto.Compilers.Relations.
+Require Import Crypto.Compilers.MapCastByDeBruijnInterp.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.MapCastByDeBruijn.
+
+Section language.
+ Context {interp_base_type_bounds : base_type -> Type}
+ (interp_op_bounds : forall src dst, op src dst -> interp_flat_type interp_base_type_bounds src -> interp_flat_type interp_base_type_bounds dst)
+ (pick_typeb : forall t, interp_base_type_bounds t -> base_type).
+ Local Notation pick_type v := (SmartFlatTypeMap pick_typeb v).
+ Context (cast_op : forall t tR (opc : op t tR) args_bs,
+ op (pick_type args_bs) (pick_type (interp_op_bounds t tR opc args_bs)))
+ (cast_backb: forall t b, interp_base_type (pick_typeb t b) -> interp_base_type t).
+ Let cast_back : forall t b, interp_flat_type interp_base_type (pick_type b) -> interp_flat_type interp_base_type t
+ := fun t b => SmartFlatTypeMapUnInterp cast_backb.
+ Context (inboundsb : forall t, interp_base_type_bounds t -> interp_base_type t -> Prop).
+ Let inbounds : forall t, interp_flat_type interp_base_type_bounds t -> interp_flat_type interp_base_type t -> Prop
+ := fun t => interp_flat_type_rel_pointwise inboundsb (t:=t).
+ Context (interp_op_bounds_correct
+ : forall t tR opc bs
+ (v : interp_flat_type interp_base_type t)
+ (H : inbounds t bs v),
+ inbounds tR (interp_op_bounds t tR opc bs) (interp_op t tR opc v))
+ (pull_cast_back
+ : forall t tR opc bs
+ (v : interp_flat_type interp_base_type (pick_type bs))
+ (H : inbounds t bs (cast_back t bs v)),
+ interp_op t tR opc (cast_back t bs v)
+ =
+ cast_back _ _ (interp_op _ _ (cast_op _ _ opc bs) v)).
+
+ Local Notation MapCast
+ := (@MapCast interp_base_type_bounds interp_op_bounds pick_typeb cast_op).
+
+ Lemma MapCastCorrect
+ {t} (e : Expr base_type op t)
+ (Hwf : Wf e)
+ (input_bounds : interp_flat_type interp_base_type_bounds (domain t))
+ : forall {b} e' (He':MapCast e input_bounds = Some (existT _ b e'))
+ v v' (Hv : @inbounds _ input_bounds v /\ cast_back _ _ v' = v),
+ Interp interp_op_bounds e input_bounds = b
+ /\ @inbounds _ b (Interp interp_op e v)
+ /\ cast_back _ _ (Interp interp_op e' v') = (Interp interp_op e v).
+ Proof using Type*.
+ apply MapCastCorrect; auto using internal_base_type_dec_lb.
+ Qed.
+End language.
diff --git a/src/Compilers/Z/MapCastByDeBruijnWf.v b/src/Compilers/Z/MapCastByDeBruijnWf.v
new file mode 100644
index 000000000..cce16e5fe
--- /dev/null
+++ b/src/Compilers/Z/MapCastByDeBruijnWf.v
@@ -0,0 +1,56 @@
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.Wf.
+Require Import Crypto.Compilers.SmartMap.
+Require Import Crypto.Compilers.Relations.
+Require Import Crypto.Compilers.MapCastByDeBruijnWf.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.MapCastByDeBruijn.
+
+Section language.
+ Context {interp_base_type_bounds : base_type -> Type}
+ (interp_op_bounds : forall src dst, op src dst -> interp_flat_type interp_base_type_bounds src -> interp_flat_type interp_base_type_bounds dst)
+ (pick_typeb : forall t, interp_base_type_bounds t -> base_type).
+ Local Notation pick_type v := (SmartFlatTypeMap pick_typeb v).
+ Context (cast_op : forall t tR (opc : op t tR) args_bs,
+ op (pick_type args_bs) (pick_type (interp_op_bounds t tR opc args_bs)))
+ (cast_backb: forall t b, interp_base_type (pick_typeb t b) -> interp_base_type t).
+ Let cast_back : forall t b, interp_flat_type interp_base_type (pick_type b) -> interp_flat_type interp_base_type t
+ := fun t b => SmartFlatTypeMapUnInterp cast_backb.
+ Context (inboundsb : forall t, interp_base_type_bounds t -> interp_base_type t -> Prop).
+ Let inbounds : forall t, interp_flat_type interp_base_type_bounds t -> interp_flat_type interp_base_type t -> Prop
+ := fun t => interp_flat_type_rel_pointwise inboundsb (t:=t).
+ Context (interp_op_bounds_correct
+ : forall t tR opc bs
+ (v : interp_flat_type interp_base_type t)
+ (H : inbounds t bs v),
+ inbounds tR (interp_op_bounds t tR opc bs) (interp_op t tR opc v))
+ (pull_cast_back
+ : forall t tR opc bs
+ (v : interp_flat_type interp_base_type (pick_type bs))
+ (H : inbounds t bs (cast_back t bs v)),
+ interp_op t tR opc (cast_back t bs v)
+ =
+ cast_back _ _ (interp_op _ _ (cast_op _ _ opc bs) v)).
+
+ Local Notation MapCast
+ := (@MapCast interp_base_type_bounds interp_op_bounds pick_typeb cast_op).
+
+ Definition Wf_MapCast
+ {t} (e : Expr base_type op t)
+ (input_bounds : interp_flat_type interp_base_type_bounds (domain t))
+ {b} e' (He' : MapCast e input_bounds = Some (existT _ b e'))
+ (Hwf : Wf e)
+ : Wf e'
+ := @Wf_MapCast
+ _ _ _ internal_base_type_dec_bl internal_base_type_dec_lb _ _ _ _ _
+ t e input_bounds b e' He' Hwf.
+ Definition Wf_MapCast_arrow
+ {s d} (e : Expr base_type op (Arrow s d))
+ (input_bounds : interp_flat_type interp_base_type_bounds s)
+ {b} e' (He' : MapCast e input_bounds = Some (existT _ b e'))
+ (Hwf : Wf e)
+ : Wf e'
+ := @Wf_MapCast_arrow
+ _ _ _ internal_base_type_dec_bl internal_base_type_dec_lb _ _ _ _ _
+ s d e input_bounds b e' He' Hwf.
+End language.
diff --git a/src/Compilers/Z/OpInversion.v b/src/Compilers/Z/OpInversion.v
new file mode 100644
index 000000000..58b00c538
--- /dev/null
+++ b/src/Compilers/Z/OpInversion.v
@@ -0,0 +1,28 @@
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.TypeInversion.
+Require Import Crypto.Compilers.Z.Syntax.
+
+Ltac invert_one_op e :=
+ preinvert_one_type e;
+ intros ? e;
+ destruct e;
+ try exact I.
+
+Ltac invert_op_step :=
+ match goal with
+ | [ e : op _ (Tbase _) |- _ ] => invert_one_op e
+ | [ e : op _ (Prod _ _) |- _ ] => invert_one_op e
+ | [ e : op _ Unit |- _ ] => invert_one_op e
+ end.
+
+Ltac invert_op := repeat invert_op_step.
+
+Ltac invert_match_op_step :=
+ match goal with
+ | [ |- appcontext[match ?e with OpConst _ _ => _ | _ => _ end] ]
+ => invert_one_op e
+ | [ H : appcontext[match ?e with OpConst _ _ => _ | _ => _ end] |- _ ]
+ => invert_one_op e
+ end.
+
+Ltac invert_match_op := repeat invert_match_op_step.
diff --git a/src/Compilers/Z/Reify.v b/src/Compilers/Z/Reify.v
new file mode 100644
index 000000000..44ac44a03
--- /dev/null
+++ b/src/Compilers/Z/Reify.v
@@ -0,0 +1,50 @@
+Require Import Coq.ZArith.ZArith.
+Require Import Crypto.Compilers.InputSyntax.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Compilers.Z.Syntax.Equality.
+Require Import Crypto.Compilers.Z.Syntax.Util.
+Require Import Crypto.Compilers.WfReflective.
+Require Import Crypto.Compilers.Reify.
+Require Import Crypto.Compilers.Inline.
+Require Import Crypto.Compilers.InlineInterp.
+Require Import Crypto.Compilers.Linearize.
+Require Import Crypto.Compilers.LinearizeInterp.
+Require Import Crypto.Compilers.Eta.
+Require Import Crypto.Compilers.EtaInterp.
+
+Ltac base_reify_op op op_head extra ::=
+ lazymatch op_head with
+ | @Z.add => constr:(reify_op op op_head 2 (Add TZ TZ TZ))
+ | @Z.mul => constr:(reify_op op op_head 2 (Mul TZ TZ TZ))
+ | @Z.sub => constr:(reify_op op op_head 2 (Sub TZ TZ TZ))
+ | @Z.shiftl => constr:(reify_op op op_head 2 (Shl TZ TZ TZ))
+ | @Z.shiftr => constr:(reify_op op op_head 2 (Shr TZ TZ TZ))
+ | @Z.land => constr:(reify_op op op_head 2 (Land TZ TZ TZ))
+ | @Z.lor => constr:(reify_op op op_head 2 (Lor TZ TZ TZ))
+ | @Z.opp => constr:(reify_op op op_head 1 (Opp TZ TZ))
+ end.
+Ltac base_reify_type T ::=
+ lazymatch T with
+ | Z => TZ
+ end.
+Ltac Reify' e := Compilers.Reify.Reify' base_type interp_base_type op e.
+Ltac Reify e :=
+ let v := Compilers.Reify.Reify base_type interp_base_type op make_const e in
+ constr:(ExprEta v).
+Ltac prove_ExprEta_Compile_correct :=
+ fun _
+ => intros;
+ rewrite ?InterpExprEta;
+ prove_compile_correct_using ltac:(fun _ => apply make_const_correct) ().
+
+Ltac Reify_rhs :=
+ Compilers.Reify.Reify_rhs_gen Reify prove_ExprEta_Compile_correct interp_op ltac:(fun tac => tac ()).
+
+Ltac prereify_context_variables :=
+ Compilers.Reify.prereify_context_variables interp_base_type.
+Ltac reify_context_variable :=
+ Compilers.Reify.reify_context_variable base_type interp_base_type op.
+Ltac lazy_reify_context_variable :=
+ Compilers.Reify.lazy_reify_context_variable base_type interp_base_type op.
+Ltac reify_context_variables :=
+ Compilers.Reify.reify_context_variables base_type interp_base_type op.
diff --git a/src/Compilers/Z/Syntax.v b/src/Compilers/Z/Syntax.v
new file mode 100644
index 000000000..754b3fb5a
--- /dev/null
+++ b/src/Compilers/Z/Syntax.v
@@ -0,0 +1,84 @@
+(** * PHOAS Syntax for expression trees on ℤ *)
+Require Import Coq.ZArith.ZArith.
+Require Import Bedrock.Word.
+Require Import Crypto.Compilers.SmartMap.
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.TypeUtil.
+Require Import Crypto.Util.FixedWordSizes.
+Require Import Crypto.Util.Option.
+Require Import Crypto.Util.NatUtil. (* for nat_beq for equality schemes *)
+Export Syntax.Notations.
+
+Local Set Boolean Equality Schemes.
+Local Set Decidable Equality Schemes.
+Inductive base_type := TZ | TWord (logsz : nat).
+
+Local Notation tZ := (Tbase TZ).
+Local Notation tWord logsz := (Tbase (TWord logsz)).
+
+Inductive op : flat_type base_type -> flat_type base_type -> Type :=
+| OpConst {T} (z : Z) : op Unit (Tbase T)
+| Add T1 T2 Tout : op (Tbase T1 * Tbase T2) (Tbase Tout)
+| Sub T1 T2 Tout : op (Tbase T1 * Tbase T2) (Tbase Tout)
+| Mul T1 T2 Tout : op (Tbase T1 * Tbase T2) (Tbase Tout)
+| Shl T1 T2 Tout : op (Tbase T1 * Tbase T2) (Tbase Tout)
+| Shr T1 T2 Tout : op (Tbase T1 * Tbase T2) (Tbase Tout)
+| Land T1 T2 Tout : op (Tbase T1 * Tbase T2) (Tbase Tout)
+| Lor T1 T2 Tout : op (Tbase T1 * Tbase T2) (Tbase Tout)
+| Opp T Tout : op (Tbase T) (Tbase Tout)
+.
+
+Definition interp_base_type (v : base_type) : Type :=
+ match v with
+ | TZ => Z
+ | TWord logsz => wordT logsz
+ end.
+
+Definition interpToZ {t} : interp_base_type t -> Z
+ := match t with
+ | TZ => fun x => x
+ | TWord _ => wordToZ
+ end.
+Definition ZToInterp {t} : Z -> interp_base_type t
+ := match t return Z -> interp_base_type t with
+ | TZ => fun x => x
+ | TWord _ => ZToWord
+ end.
+Definition cast_const {t1 t2} (v : interp_base_type t1) : interp_base_type t2
+ := ZToInterp (interpToZ v).
+
+Local Notation eta x := (fst x, snd x).
+Local Notation eta3 x := (eta (fst x), snd x).
+Local Notation eta4 x := (eta3 (fst x), snd x).
+
+Definition lift_op {src dst}
+ (srcv:=SmartValf (fun _ => base_type) (fun t => t) src)
+ (dstv:=SmartValf (fun _ => base_type) (fun t => t) dst)
+ (ff:=fun t0 (v : interp_flat_type _ t0) t => SmartFlatTypeMap (var':=fun _ => base_type) (fun _ _ => t) v)
+ (srcf:=ff src srcv) (dstf:=ff dst dstv)
+ (srcZ:=srcf TZ) (dstZ:=dstf TZ)
+ (opZ : interp_flat_type interp_base_type srcZ -> interp_flat_type interp_base_type dstZ)
+ : interp_flat_type interp_base_type src
+ -> interp_flat_type interp_base_type dst
+ := fun xy
+ => SmartFlatTypeMapUnInterp
+ (fun _ _ => cast_const)
+ (opZ (SmartFlatTypeMapInterp2 (fun _ _ => cast_const) _ xy)).
+
+Definition Zinterp_op src dst (f : op src dst)
+ (asZ := fun t0 => SmartFlatTypeMap (var':=fun _ => base_type) (fun _ _ => TZ) (SmartValf (fun _ => base_type) (fun t => t) t0))
+ : interp_flat_type interp_base_type (asZ src) -> interp_flat_type interp_base_type (asZ dst)
+ := match f in op src dst return interp_flat_type interp_base_type (asZ src) -> interp_flat_type interp_base_type (asZ dst) with
+ | OpConst _ v => fun _ => cast_const (t1:=TZ) v
+ | Add _ _ _ => fun xy => fst xy + snd xy
+ | Sub _ _ _ => fun xy => fst xy - snd xy
+ | Mul _ _ _ => fun xy => fst xy * snd xy
+ | Shl _ _ _ => fun xy => Z.shiftl (fst xy) (snd xy)
+ | Shr _ _ _ => fun xy => Z.shiftr (fst xy) (snd xy)
+ | Land _ _ _ => fun xy => Z.land (fst xy) (snd xy)
+ | Lor _ _ _ => fun xy => Z.lor (fst xy) (snd xy)
+ | Opp _ _ => fun x => Z.opp x
+ end%Z.
+
+Definition interp_op src dst (f : op src dst) : interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst
+ := lift_op (Zinterp_op src dst f).
diff --git a/src/Compilers/Z/Syntax/Equality.v b/src/Compilers/Z/Syntax/Equality.v
new file mode 100644
index 000000000..10cd287be
--- /dev/null
+++ b/src/Compilers/Z/Syntax/Equality.v
@@ -0,0 +1,176 @@
+Require Import Coq.ZArith.ZArith.
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.TypeInversion.
+Require Import Crypto.Compilers.Equality.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Util.PartiallyReifiedProp.
+Require Import Crypto.Util.HProp.
+Require Import Crypto.Util.Tactics.BreakMatch.
+Require Import Crypto.Util.Tactics.DestructHead.
+Require Import Crypto.Util.FixedWordSizesEquality.
+Require Import Crypto.Util.NatUtil.
+
+Global Instance dec_eq_base_type : DecidableRel (@eq base_type)
+ := base_type_eq_dec.
+Global Instance dec_eq_flat_type : DecidableRel (@eq (flat_type base_type)) := _.
+Global Instance dec_eq_type : DecidableRel (@eq (type base_type)) := _.
+
+Definition base_type_eq_semidec_transparent (t1 t2 : base_type)
+ : option (t1 = t2)
+ := match base_type_eq_dec t1 t2 with
+ | left pf => Some pf
+ | right _ => None
+ end.
+Lemma base_type_eq_semidec_is_dec t1 t2 : base_type_eq_semidec_transparent t1 t2 = None -> t1 <> t2.
+Proof.
+ unfold base_type_eq_semidec_transparent; break_match; congruence.
+Qed.
+
+Definition op_beq_hetero {t1 tR t1' tR'} (f : op t1 tR) (g : op t1' tR') : bool
+ := match f, g return bool with
+ | OpConst T1 v, OpConst T2 v'
+ => base_type_beq T1 T2 && Z.eqb v v'
+ | Add T1 T2 Tout, Add T1' T2' Tout'
+ | Sub T1 T2 Tout, Sub T1' T2' Tout'
+ | Mul T1 T2 Tout, Mul T1' T2' Tout'
+ | Shl T1 T2 Tout, Shl T1' T2' Tout'
+ | Shr T1 T2 Tout, Shr T1' T2' Tout'
+ | Land T1 T2 Tout, Land T1' T2' Tout'
+ | Lor T1 T2 Tout, Lor T1' T2' Tout'
+ => base_type_beq T1 T1' && base_type_beq T2 T2' && base_type_beq Tout Tout'
+ | Opp Tin Tout, Opp Tin' Tout'
+ => base_type_beq Tin Tin' && base_type_beq Tout Tout'
+ | OpConst _ _, _
+ | Add _ _ _, _
+ | Sub _ _ _, _
+ | Mul _ _ _, _
+ | Shl _ _ _, _
+ | Shr _ _ _, _
+ | Land _ _ _, _
+ | Lor _ _ _, _
+ | Opp _ _, _
+ => false
+ end%bool.
+
+Definition op_beq t1 tR (f g : op t1 tR) : bool
+ := Eval cbv [op_beq_hetero] in op_beq_hetero f g.
+
+Definition op_beq_hetero_type_eq {t1 tR t1' tR'} f g : to_prop (@op_beq_hetero t1 tR t1' tR' f g) -> t1 = t1' /\ tR = tR'.
+Proof.
+ destruct f, g;
+ repeat match goal with
+ | _ => progress unfold op_beq_hetero in *
+ | _ => simpl; intro; exfalso; assumption
+ | _ => solve [ repeat constructor ]
+ | [ |- context[reified_Prop_of_bool ?b] ]
+ => let H := fresh in destruct (Sumbool.sumbool_of_bool b) as [H|H]; rewrite H
+ | [ H : nat_beq _ _ = true |- _ ] => apply internal_nat_dec_bl in H; subst
+ | [ H : base_type_beq _ _ = true |- _ ] => apply internal_base_type_dec_bl in H; subst
+ | [ H : wordT_beq_hetero _ _ = true |- _ ] => apply wordT_beq_bl in H; subst
+ | [ H : wordT_beq_hetero _ _ = true |- _ ] => apply wordT_beq_hetero_bl in H; destruct H; subst
+ | [ H : andb ?x ?y = true |- _ ]
+ => assert (x = true /\ y = true) by (destruct x, y; simpl in *; repeat constructor; exfalso; clear -H; abstract congruence);
+ clear H
+ | [ H : and _ _ |- _ ] => destruct H
+ | [ H : false = true |- _ ] => exfalso; clear -H; abstract congruence
+ | [ H : true = false |- _ ] => exfalso; clear -H; abstract congruence
+ | _ => progress break_match_hyps
+ end.
+Defined.
+
+Definition op_beq_hetero_type_eqs {t1 tR t1' tR'} f g : to_prop (@op_beq_hetero t1 tR t1' tR' f g) -> t1 = t1'
+ := fun H => let (p, q) := @op_beq_hetero_type_eq t1 tR t1' tR' f g H in p.
+Definition op_beq_hetero_type_eqd {t1 tR t1' tR'} f g : to_prop (@op_beq_hetero t1 tR t1' tR' f g) -> tR = tR'
+ := fun H => let (p, q) := @op_beq_hetero_type_eq t1 tR t1' tR' f g H in q.
+
+Definition op_beq_hetero_eq {t1 tR t1' tR'} f g
+ : forall pf : to_prop (@op_beq_hetero t1 tR t1' tR' f g),
+ eq_rect
+ _ (fun src => op src tR')
+ (eq_rect _ (fun dst => op t1 dst) f _ (op_beq_hetero_type_eqd f g pf))
+ _ (op_beq_hetero_type_eqs f g pf)
+ = g.
+Proof.
+ destruct f, g;
+ repeat match goal with
+ | _ => solve [ intros [] ]
+ | _ => reflexivity
+ | [ H : False |- _ ] => exfalso; assumption
+ | _ => intro
+ | [ |- context[op_beq_hetero_type_eqd ?f ?g ?pf] ]
+ => generalize (op_beq_hetero_type_eqd f g pf), (op_beq_hetero_type_eqs f g pf)
+ | _ => intro
+ | _ => progress eliminate_hprop_eq
+ | _ => progress inversion_flat_type
+ | _ => progress unfold op_beq_hetero in *
+ | _ => progress simpl in *
+ | [ H : context[andb ?x ?y] |- _ ]
+ => destruct x eqn:?, y eqn:?; simpl in H
+ | [ H : Z.eqb _ _ = true |- _ ] => apply Z.eqb_eq in H
+ | [ H : to_prop (reified_Prop_of_bool ?b) |- _ ] => destruct b eqn:?; compute in H
+ | _ => progress subst
+ | _ => progress break_match_hyps
+ | [ H : wordT_beq_hetero _ _ = true |- _ ] => apply wordT_beq_bl in H; subst
+ | [ H : wordT_beq_hetero _ _ = true |- _ ] => apply wordT_beq_hetero_bl in H; destruct H; subst
+ | _ => congruence
+ end.
+Qed.
+
+Lemma op_beq_bl : forall t1 tR x y, to_prop (op_beq t1 tR x y) -> x = y.
+Proof.
+ intros ?? f g H.
+ pose proof (op_beq_hetero_eq f g H) as H'.
+ generalize dependent (op_beq_hetero_type_eqd f g H).
+ generalize dependent (op_beq_hetero_type_eqs f g H).
+ intros; eliminate_hprop_eq; simpl in *; assumption.
+Qed.
+
+Section encode_decode.
+ Definition base_type_code (t1 t2 : base_type) : Prop
+ := match t1, t2 with
+ | TZ, TZ => True
+ | TWord s1, TWord s2 => s1 = s2
+ | TZ, _
+ | TWord _, _
+ => False
+ end.
+
+ Definition base_type_encode (x y : base_type) : x = y -> base_type_code x y.
+ Proof. intro p; destruct p, x; repeat constructor. Defined.
+
+ Definition base_type_decode (x y : base_type) : base_type_code x y -> x = y.
+ Proof.
+ destruct x, y; simpl in *; intro H;
+ try first [ apply f_equal; assumption
+ | exfalso; assumption
+ | reflexivity
+ | apply f_equal2; destruct H; assumption ].
+ Defined.
+ Definition path_base_type_rect {x y : base_type} (Q : x = y -> Type)
+ (f : forall p, Q (base_type_decode x y p))
+ : forall p, Q p.
+ Proof. intro p; specialize (f (base_type_encode x y p)); destruct x, p; exact f. Defined.
+End encode_decode.
+
+Ltac induction_type_in_using H rect :=
+ induction H as [H] using (rect _ _);
+ cbv [base_type_code] in H;
+ let H1 := fresh H in
+ let H2 := fresh H in
+ try lazymatch type of H with
+ | False => exfalso; exact H
+ | True => destruct H
+ end.
+Ltac inversion_base_type_step :=
+ lazymatch goal with
+ | [ H : _ = TWord _ |- _ ]
+ => induction_type_in_using H @path_base_type_rect
+ | [ H : TWord _ = _ |- _ ]
+ => induction_type_in_using H @path_base_type_rect
+ | [ H : _ = TZ |- _ ]
+ => induction_type_in_using H @path_base_type_rect
+ | [ H : TZ = _ |- _ ]
+ => induction_type_in_using H @path_base_type_rect
+ end.
+Ltac inversion_base_type := repeat inversion_base_type_step.
diff --git a/src/Compilers/Z/Syntax/Util.v b/src/Compilers/Z/Syntax/Util.v
new file mode 100644
index 000000000..6cf54a99c
--- /dev/null
+++ b/src/Compilers/Z/Syntax/Util.v
@@ -0,0 +1,170 @@
+Require Import Crypto.Compilers.Syntax.
+Require Import Crypto.Compilers.SmartMap.
+Require Import Crypto.Compilers.Wf.
+Require Import Crypto.Compilers.TypeUtil.
+Require Import Crypto.Compilers.TypeInversion.
+Require Import Crypto.Compilers.Z.Syntax.
+Require Import Crypto.Util.FixedWordSizesEquality.
+Require Import Crypto.Util.NatUtil.
+Require Import Crypto.Util.HProp.
+Require Import Crypto.Util.Tactics.BreakMatch.
+Require Import Crypto.Util.Tactics.DestructHead.
+Require Import Crypto.Util.Notations.
+
+Definition make_const t : interp_base_type t -> op Unit (Tbase t)
+ := fun v => OpConst (cast_const (t2:=TZ) v).
+Definition is_const s d (v : op s d) : bool
+ := match v with OpConst _ _ => true | _ => false end.
+Arguments is_const [s d] v.
+
+Definition cast_back_flat_const {var t f V}
+ (v : interp_flat_type interp_base_type (@SmartFlatTypeMap base_type var f t V))
+ : interp_flat_type interp_base_type t
+ := @SmartFlatTypeMapUnInterp
+ _ var interp_base_type interp_base_type
+ f (fun _ _ => cast_const)
+ t V v.
+
+Definition cast_flat_const {var t f V}
+ (v : interp_flat_type interp_base_type t)
+ : interp_flat_type interp_base_type (@SmartFlatTypeMap base_type var f t V)
+ := @SmartFlatTypeMapInterp2
+ _ var interp_base_type interp_base_type
+ f (fun _ _ => cast_const)
+ t V v.
+
+Definition base_type_leb (v1 v2 : base_type) : bool
+ := match v1, v2 with
+ | _, TZ => true
+ | TZ, _ => false
+ | TWord logsz1, TWord logsz2 => Compare_dec.leb logsz1 logsz2
+ end.
+
+Definition base_type_min := base_type_min base_type_leb.
+Definition base_type_max := base_type_max base_type_leb.
+Global Arguments base_type_min !_ !_ / .
+Global Arguments base_type_max !_ !_ / .
+Global Arguments TypeUtil.base_type_min _ _ _ / _.
+Global Arguments TypeUtil.base_type_max _ _ _ / _.
+
+Definition genericize_op {var' src dst} (opc : op src dst) {f}
+ : forall {vs vd}, op (@SmartFlatTypeMap _ var' f src vs) (@SmartFlatTypeMap _ var' f dst vd)
+ := match opc with
+ | OpConst _ z => fun _ _ => OpConst z
+ | Add _ _ _ => fun _ _ => Add _ _ _
+ | Sub _ _ _ => fun _ _ => Sub _ _ _
+ | Mul _ _ _ => fun _ _ => Mul _ _ _
+ | Shl _ _ _ => fun _ _ => Shl _ _ _
+ | Shr _ _ _ => fun _ _ => Shr _ _ _
+ | Land _ _ _ => fun _ _ => Land _ _ _
+ | Lor _ _ _ => fun _ _ => Lor _ _ _
+ | Opp _ _ => fun _ _ => Opp _ _
+ end.
+
+Lemma cast_const_id {t} v
+ : @cast_const t t v = v.
+Proof.
+ destruct t; simpl; trivial.
+ rewrite ZToWord_wordToZ; reflexivity.
+Qed.
+
+Lemma cast_const_idempotent {a b c} v
+ : base_type_min b (base_type_min a c) = base_type_min a c
+ -> @cast_const b c (@cast_const a b v) = @cast_const a c v.
+Proof.
+ repeat first [ reflexivity
+ | congruence
+ | progress destruct_head' base_type
+ | progress simpl
+ | progress break_match
+ | progress subst
+ | intro
+ | match goal with
+ | [ H : ?leb _ _ = true |- _ ] => apply Compare_dec.leb_complete in H
+ | [ H : ?leb _ _ = false |- _ ] => apply Compare_dec.leb_iff_conv in H
+ | [ H : TWord _ = TWord _ |- _ ] => inversion H; clear H
+ end
+ | rewrite ZToWord_wordToZ_ZToWord by omega *
+ | rewrite wordToZ_ZToWord_wordToZ by omega * ].
+Qed.
+
+Lemma make_const_correct : forall T v, interp_op Unit (Tbase T) (make_const T v) tt = v.
+Proof.
+ destruct T; cbv -[FixedWordSizes.ZToWord FixedWordSizes.wordToZ FixedWordSizes.wordT];
+ intro; rewrite ?ZToWord_wordToZ; reflexivity.
+Qed.
+
+Local Notation iffT A B := ((A -> B) * (B -> A))%type (only parsing).
+
+Section unzify.
+ Context {var'} {f : forall t : base_type, var' t -> base_type}.
+ Let asZ := fun t => SmartFlatTypeMap
+ (fun _ _ => TZ)
+ (SmartValf (fun _ => base_type) (fun t => t) t).
+ Definition unzify_op_helper_step
+ (unzify_op_helper
+ : forall {t : flat_type base_type}
+ {vs : interp_flat_type var' t},
+ iffT (interp_flat_type
+ interp_base_type
+ (asZ t))
+ (interp_flat_type
+ interp_base_type
+ (asZ (SmartFlatTypeMap f vs))))
+ {t : flat_type base_type}
+ : forall {vs : interp_flat_type var' t},
+ iffT (interp_flat_type
+ interp_base_type
+ (asZ t))
+ (interp_flat_type
+ interp_base_type
+ (asZ (SmartFlatTypeMap f vs)))
+ := match t with
+ | Tbase T => fun _ => (fun x => x, fun x => x)
+ | Unit => fun _ => (fun x => x, fun x => x)
+ | Prod A B
+ => fun (vs : interp_flat_type _ A * interp_flat_type _ B)
+ => let f1 := @unzify_op_helper A (fst vs) in
+ let f2 := @unzify_op_helper B (snd vs) in
+ ((fun x : interp_flat_type _ (asZ A) * interp_flat_type _ (asZ B)
+ => (fst f1 (fst x), fst f2 (snd x))),
+ (fun x : interp_flat_type _ (asZ (SmartFlatTypeMap f (fst vs)))
+ * interp_flat_type _ (asZ (SmartFlatTypeMap f (snd vs)))
+ => (snd f1 (fst x), snd f2 (snd x))))
+ end.
+ Fixpoint unzify_op_helper {t vs}
+ := @unzify_op_helper_step (@unzify_op_helper) t vs.
+
+ Definition unzify_op
+ {src dst : flat_type base_type}
+ {vs : interp_flat_type var' src} {vd : interp_flat_type var' dst}
+ (F : interp_flat_type interp_base_type (asZ src) -> interp_flat_type interp_base_type (asZ dst))
+ (x : interp_flat_type interp_base_type (asZ (SmartFlatTypeMap f vs)))
+ : interp_flat_type interp_base_type (asZ (SmartFlatTypeMap f vd))
+ := fst unzify_op_helper (F (snd unzify_op_helper x)).
+End unzify.
+
+Arguments unzify_op_helper_step _ _ _ !_ _ / .
+Arguments unzify_op_helper _ _ !_ _ / .
+
+Lemma Zinterp_op_genericize_op {var' src dst opc f vs vd}
+ : Zinterp_op _ _ (@genericize_op var' src dst opc f vs vd)
+ = unzify_op (Zinterp_op _ _ opc).
+Proof.
+ destruct opc; unfold unzify_op; reflexivity.
+Qed.
+
+Lemma lift_op_prod_dst {src dstA dstB}
+ {f : _ -> interp_flat_type _ (SmartFlatTypeMap _ (SmartValf _ _ _)) * interp_flat_type _ (SmartFlatTypeMap _ (SmartValf _ _ _))}
+ {x}
+ : @lift_op src (Prod dstA dstB) f x
+ = (@lift_op src dstA (fun y => fst (f y)) x, @lift_op src dstB (fun y => snd (f y)) x).
+Proof. reflexivity. Qed.
+
+Lemma cast_back_flat_const_prod {var A B f} {V : _ * _}
+ (v : interp_flat_type interp_base_type (@SmartFlatTypeMap base_type var f A (fst V))
+ * interp_flat_type interp_base_type (@SmartFlatTypeMap base_type var f B (snd V)))
+ : @cast_back_flat_const var (Prod A B) f V v
+ = (@cast_back_flat_const var A f (fst V) (fst v),
+ @cast_back_flat_const var B f (snd V) (snd v)).
+Proof. reflexivity. Qed.